From d0257b0a47ad998e01715e9bc6ba612b834765f1 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Thu, 11 Jun 2020 14:47:52 +0100 Subject: Working on proof. --- src/translation/HTLgen.v | 7 +- src/translation/HTLgenproof.v | 152 ++++++++++++++++++++++++++++++++---------- src/translation/HTLgenspec.v | 18 ++++- src/verilog/HTL.v | 21 ++++-- src/verilog/Value.v | 5 ++ 5 files changed, 156 insertions(+), 47 deletions(-) diff --git a/src/translation/HTLgen.v b/src/translation/HTLgen.v index b573b06..d5a8af2 100644 --- a/src/translation/HTLgen.v +++ b/src/translation/HTLgen.v @@ -398,9 +398,9 @@ Lemma create_arr_state_incr: (st_controllogic s)). Proof. constructor; simpl; auto with htlh. Qed. -Definition create_arr (sz : nat) (ln : nat) : mon reg := +Definition create_arr (sz : nat) (ln : nat) : mon (reg * nat) := fun s => let r := s.(st_freshreg) in - OK r (mkstate + OK (r, ln) (mkstate s.(st_st) (Pos.succ r) (st_freshstate s) @@ -413,7 +413,7 @@ Definition create_arr (sz : nat) (ln : nat) : mon reg := Definition transf_module (f: function) : mon module := do fin <- create_reg 1; do rtrn <- create_reg 32; - do stack <- create_arr 32 (Z.to_nat (f.(fn_stacksize) / 4)); + do (stack, stack_len) <- create_arr 32 (Z.to_nat (f.(fn_stacksize) / 4)); do _ <- collectlist (transf_instr fin rtrn stack) (Maps.PTree.elements f.(RTL.fn_code)); do start <- create_reg 1; do rst <- create_reg 1; @@ -426,6 +426,7 @@ Definition transf_module (f: function) : mon module := f.(fn_entrypoint) current_state.(st_st) stack + stack_len fin rtrn). diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index 523c86c..1f9e368 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -42,44 +42,45 @@ Definition state_st_wf (m : HTL.module) (s : HTL.state) := asa!(m.(HTL.mod_st)) = Some (posToValue 32 st). Hint Unfold state_st_wf : htlproof. -Inductive match_stacks : list RTL.stackframe -> list HTL.stackframe -> Prop := -| match_stacks_nil : - match_stacks nil nil -| match_stacks_cons : - forall cs lr r f sp pc rs m assoc - (TF : tr_module f m) - (ST: match_stacks cs lr) - (MA: match_assocmaps f rs assoc), - match_stacks (RTL.Stackframe r f sp pc rs :: cs) - (HTL.Stackframe r m pc assoc :: lr). - -Inductive match_arrs (m : HTL.module) : RTL.function -> mem -> AssocMap.t (list value) -> Prop := -| match_arr : forall mem asa stack sp f sz, - sz = f.(RTL.fn_stacksize) -> +Inductive match_arrs (m : HTL.module) (f : RTL.function) (sp : Values.val) (mem : mem) : + AssocMap.t (list value) -> Prop := +| match_arr : forall asa stack, + m.(HTL.mod_stk_len) = Z.to_nat (f.(RTL.fn_stacksize) / 4) -> asa ! (m.(HTL.mod_stk)) = Some stack -> (forall ptr, - 0 <= ptr < sz -> - nth_error stack (Z.to_nat ptr) = - (Coquplib.Option.bind (Mem.loadv AST.Mint32 mem + 0 <= ptr < Z.of_nat m.(HTL.mod_stk_len) -> + opt_val_value_lessdef (Mem.loadv AST.Mint32 mem (Values.Val.offset_ptr sp (Integers.Ptrofs.repr ptr))) - valToValue)) -> - match_arrs m f mem asa. + (nth (Z.to_nat ptr) stack (ZToValue 32 0))) -> + match_arrs m f sp mem asa. + +Inductive match_stacks (mem : mem) : list RTL.stackframe -> list HTL.stackframe -> Prop := +| match_stacks_nil : + match_stacks mem nil nil +| match_stacks_cons : + forall cs lr r f sp pc rs m asr asa + (TF : tr_module f m) + (ST: match_stacks mem cs lr) + (MA: match_assocmaps f rs asr) + (MARR : match_arrs m f sp mem asa), + match_stacks mem (RTL.Stackframe r f sp pc rs :: cs) + (HTL.Stackframe r m pc asr asa :: lr). Inductive match_states : RTL.state -> HTL.state -> Prop := | match_state : forall asa asr sf f sp rs mem m st res (MASSOC : match_assocmaps f rs asr) (TF : tr_module f m) (WF : state_st_wf m (HTL.State res m st asr asa)) - (MS : match_stacks sf res) - (MA : match_arrs m f mem asa), + (MS : match_stacks mem sf res) + (MARR : match_arrs m f sp mem asa), match_states (RTL.State sf f sp st rs mem) (HTL.State res m st asr asa) | match_returnstate : forall - v v' stack m res - (MS : match_stacks stack res), + v v' stack mem res + (MS : match_stacks mem stack res), val_value_lessdef v v' -> - match_states (RTL.Returnstate stack v m) (HTL.Returnstate res v') + match_states (RTL.Returnstate stack v mem) (HTL.Returnstate res v') | match_initial_call : forall f m m0 (TF : tr_module f m), @@ -390,7 +391,73 @@ Section CORRECTNESS. (* assumption. *) admit. - - admit. + Ltac invert x := inversion x; subst; clear x. + + Ltac clear_obvious := + repeat match goal with + | [ H : ex _ |- _ ] => invert H + | [ H : Some _ = Some _ |- _ ] => invert H + end. + + Ltac simplify := simpl in *; clear_obvious; simpl in *; try discriminate. + + Ltac rt := + repeat match goal with + | [ _ : error _ _ = OK _ _ _ |- _ ] => discriminate + | [ _ : context[if (?x && ?y) then _ else _] |- _ ] => + let EQ1 := fresh "EQ" in + let EQ2 := fresh "EQ" in + destruct x eqn:EQ1; destruct y eqn:EQ2; simpl in * + | [ _ : context[if ?x then _ else _] |- _ ] => + let EQ := fresh "EQ" in + destruct x eqn:EQ; simpl in * + | [ H : ret _ _ = _ |- _ ] => invert H + | [ _ : context[match ?x with | _ => _ end] |- _ ] => destruct x + end. + + Ltac t := + match goal with + | [ _ : Mem.loadv _ _ ?a = Some _ |- _ ] => + let PTR := fresh "PTR" in + assert (exists b ofs, a = Values.Vptr b ofs) as PTR; + [> destruct a; simpl in *; try discriminate; + repeat eexists |] + | [ H: Values.Vptr _ _ = Values.Vptr _ _ |- _] => invert H + end. + + - destruct c, addr, args; simplify; rt; t; simplify. + + + admit. (* FIXME: Alignment *) + + + + unfold Op.eval_addressing in *. + (* destruct (Archi.ptr64); simplify; *) + (* destruct (Registers.Regmap.get r0 rs); simplify. *) + admit. + + (* eexists. split. *) + (* eapply Smallstep.plus_one. *) + (* eapply HTL.step_module; eauto. *) + (* econstructor. econstructor. econstructor. simpl. *) + (* econstructor. econstructor. econstructor. econstructor. econstructor. *) + (* econstructor. econstructor. econstructor. econstructor. econstructor. *) + (* econstructor. econstructor. econstructor. simpl. *) + + (* inversion MARR; subst; clear MARR. simpl in *. *) + (* unfold Verilog.arr_assocmap_lookup. *) + (* rewrite H3. *) + + (* admit. *) + + + unfold Op.eval_addressing in *. + (* destruct (Archi.ptr64); simplify; *) + (* destruct (Registers.Regmap.get r0 rs); *) + (* destruct (Registers.Regmap.get r1 rs); simplify; *) + (* destruct (Archi.ptr64); simplify. *) + admit. + + + admit. + - admit. - eexists. split. apply Smallstep.plus_one. @@ -453,8 +520,7 @@ Section CORRECTNESS. econstructor; simpl; trivial. constructor. - constructor. - constructor. + constructor. constructor. unfold_merge. simpl. rewrite AssocMap.gso. @@ -470,7 +536,8 @@ Section CORRECTNESS. apply AssocMap.gss. rewrite Events.E0_left. reflexivity. constructor. - assumption. + + admit. constructor. - econstructor. split. @@ -499,21 +566,38 @@ Section CORRECTNESS. apply AssocMap.gss. rewrite Events.E0_left. trivial. - constructor. assumption. simpl. inversion MASSOC. subst. + constructor. + admit. + + simpl. inversion MASSOC. subst. unfold find_assocmap, AssocMapExt.get_default. rewrite AssocMap.gso. apply H1. eapply RTL.max_reg_function_use. eauto. simpl; tauto. apply st_greater_than_res. - inversion MSTATE; subst. inversion TF; subst. econstructor. split. apply Smallstep.plus_one. - eapply HTL.step_call. + eapply HTL.step_call. simpl. constructor. apply regs_lessdef_add_greater. apply greater_than_max_func. apply init_reg_assoc_empty. assumption. unfold state_st_wf. intros. inv H1. apply AssocMap.gss. constructor. - admit. + econstructor; simpl. + reflexivity. + rewrite AssocMap.gss. reflexivity. + intros. + destruct (Mem.load AST.Mint32 m' stk + (Integers.Ptrofs.unsigned (Integers.Ptrofs.add + Integers.Ptrofs.zero + (Integers.Ptrofs.repr ptr)))) eqn:LOAD. + pose proof Mem.load_alloc_same as LOAD_ALLOC. + pose proof H as ALLOC. + eapply LOAD_ALLOC in ALLOC. + 2: { exact LOAD. } + rewrite ALLOC. + repeat constructor. + constructor. - inversion MSTATE; subst. inversion MS; subst. econstructor. split. apply Smallstep.plus_one. @@ -525,13 +609,11 @@ Section CORRECTNESS. unfold state_st_wf. intros. inv H. rewrite AssocMap.gso. rewrite AssocMap.gss. trivial. apply st_greater_than_res. - admit. - Unshelve. exact (AssocMap.empty value). exact (AssocMap.empty value). - admit. - admit. + exact (AssocMap.empty value). + exact (AssocMap.empty value). (* exact (ZToValue 32 0). *) (* exact (AssocMap.empty value). *) (* exact (AssocMap.empty value). *) diff --git a/src/translation/HTLgenspec.v b/src/translation/HTLgenspec.v index 4bf3843..89b79ac 100644 --- a/src/translation/HTLgenspec.v +++ b/src/translation/HTLgenspec.v @@ -184,14 +184,15 @@ Hint Constructors tr_code : htlspec. Inductive tr_module (f : RTL.function) : module -> Prop := tr_module_intro : - forall data control fin rtrn st stk m, + forall data control fin rtrn st stk stk_len m, (forall pc i, Maps.PTree.get pc f.(RTL.fn_code) = Some i -> - tr_code f.(RTL.fn_code) pc i data control fin rtrn st stk) -> + tr_code f.(RTL.fn_code) pc i data control fin rtrn st stk) -> + stk_len = Z.to_nat (f.(RTL.fn_stacksize) / 4) -> m = (mkmodule f.(RTL.fn_params) data control f.(RTL.fn_entrypoint) - st stk fin rtrn) -> + st stk stk_len fin rtrn) -> tr_module f m. Hint Constructors tr_module : htlspec. @@ -379,6 +380,12 @@ Proof. Qed. Hint Resolve iter_expand_instr_spec : htlspec. +Lemma create_arr_inv : forall x y z a b c d, + create_arr x y z = OK (a, b) c d -> y = b. +Proof. + inversion 1. reflexivity. +Qed. + Theorem transl_module_correct : forall f m, transl_module f = Errors.OK m -> tr_module f m. @@ -392,6 +399,11 @@ Proof. unfold transf_module in *. monadInv Heqr. + + (* TODO: We should be able to fold this into the automation. *) + pose proof (create_arr_inv _ _ _ _ _ _ _ EQ0) as STK_LEN. + rewrite <- STK_LEN. + econstructor; simpl; trivial. intros. inv_incr. diff --git a/src/verilog/HTL.v b/src/verilog/HTL.v index 2e4ef1a..82aac41 100644 --- a/src/verilog/HTL.v +++ b/src/verilog/HTL.v @@ -46,6 +46,7 @@ Record module: Type := mod_entrypoint : node; mod_st : reg; mod_stk : reg; + mod_stk_len : nat; mod_finish : reg; mod_return : reg }. @@ -60,6 +61,14 @@ Fixpoint init_regs (vl : list value) (rl : list reg) {struct rl} := | _, _ => empty_assocmap end. +Fixpoint zeroes' (acc : list value) (n : nat) : list value := + match n with + | O => acc + | S n => zeroes' ((NToValue 32 0)::acc) n + end. + +Definition zeroes : nat -> list value := zeroes' nil. + (** * Operational Semantics *) Definition genv := Globalenvs.Genv.t fundef unit. @@ -69,7 +78,8 @@ Inductive stackframe : Type := forall (res : reg) (m : module) (pc : node) - (assoc : assocmap), + (reg_assoc : assocmap) + (arr_assoc : AssocMap.t (list value)), stackframe. Inductive state : Type := @@ -125,13 +135,12 @@ Inductive step : genv -> state -> Events.trace -> state -> Prop := (State res m m.(mod_entrypoint) (AssocMap.set (mod_st m) (posToValue 32 m.(mod_entrypoint)) (init_regs args m.(mod_params))) - (AssocMap.empty (list value))) + (AssocMap.set m.(mod_stk) (zeroes m.(mod_stk_len)) (AssocMap.empty (list value)))) | step_return : - forall g m asr i r sf pc mst, + forall g m asr asa i r sf pc mst, mst = mod_st m -> - step g (Returnstate (Stackframe r m pc asr :: sf) i) Events.E0 - (State sf m pc ((asr # mst <- (posToValue 32 pc)) # r <- i) - (AssocMap.empty (list value))). + step g (Returnstate (Stackframe r m pc asr asa :: sf) i) Events.E0 + (State sf m pc ((asr # mst <- (posToValue 32 pc)) # r <- i) asa). Hint Constructors step : htl. Inductive initial_state (p: program): state -> Prop := diff --git a/src/verilog/Value.v b/src/verilog/Value.v index d527b15..b1ee353 100644 --- a/src/verilog/Value.v +++ b/src/verilog/Value.v @@ -296,6 +296,11 @@ Inductive val_value_lessdef: val -> value -> Prop := val_value_lessdef (Vint i) v' | lessdef_undef: forall v, val_value_lessdef Vundef v. +Inductive opt_val_value_lessdef: option val -> value -> Prop := +| opt_lessdef_some: + forall v v', val_value_lessdef v v' -> opt_val_value_lessdef (Some v) v' +| opt_lessdef_none: forall v, opt_val_value_lessdef None v. + Lemma valueToZ_ZToValue : forall n z, (- Z.of_nat (2 ^ n) <= z < Z.of_nat (2 ^ n))%Z -> -- cgit From 088a554043e3d4b8b8b424dbda9a136e3f4571e5 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Thu, 11 Jun 2020 22:50:01 +0100 Subject: Rough outline of stack address proof --- src/translation/HTLgen.v | 12 ++--- src/translation/HTLgenproof.v | 123 +++++++++++++++++++++++++++++++++--------- src/verilog/Verilog.v | 8 +-- 3 files changed, 109 insertions(+), 34 deletions(-) diff --git a/src/translation/HTLgen.v b/src/translation/HTLgen.v index d5a8af2..cba2940 100644 --- a/src/translation/HTLgen.v +++ b/src/translation/HTLgen.v @@ -314,21 +314,21 @@ Definition add_branch_instr (e: expr) (n n1 n2: node) : mon unit := Definition translate_arr_access (mem : AST.memory_chunk) (addr : Op.addressing) (args : list reg) (stack : reg) : mon expr := - match addr, args with (* TODO: We should be more methodical here; what are the possibilities?*) - | Op.Aindexed off, r1::nil => ret (Vvari stack (boplitz Vadd r1 off)) (* FIXME: Cannot guarantee alignment *) - | Op.Ascaled scale offset, r1::nil => + match mem, addr, args with (* TODO: We should be more methodical here; what are the possibilities?*) + | Mint32, Op.Aindexed off, r1::nil => ret (Vvari stack (boplitz Vadd r1 off)) (* FIXME: Cannot guarantee alignment *) + | Mint32, Op.Ascaled scale offset, r1::nil => if ((Z.eqb (Z.modulo scale 4) 0) && (Z.eqb (Z.modulo offset 4) 0)) then ret (Vvari stack (Vbinop Vadd (boplitz Vmul r1 (scale / 4)) (Vlit (ZToValue 32 (offset / 4))))) else error (Errors.msg "Veriloggen: translate_arr_access address misaligned") - | Op.Aindexed2scaled scale offset, r1::r2::nil => (* Typical for dynamic array addressing *) + | Mint32, Op.Aindexed2scaled scale offset, r1::r2::nil => (* Typical for dynamic array addressing *) if ((Z.eqb (Z.modulo scale 4) 0) && (Z.eqb (Z.modulo offset 4) 0)) then ret (Vvari stack (Vbinop Vadd (boplitz Vadd r1 (offset / 4)) (boplitz Vmul r2 (scale / 4)))) else error (Errors.msg "Veriloggen: translate_arr_access address misaligned") - | Op.Ainstack a, nil => (* We need to be sure that the base address is aligned *) + | Mint32, Op.Ainstack a, nil => (* We need to be sure that the base address is aligned *) let a := Integers.Ptrofs.unsigned a in (* FIXME: Assuming stack offsets are +ve; is this ok? *) if (Z.eq_dec (Z.modulo a 4) 0) then ret (Vvari stack (Vlit (ZToValue 32 (a / 4)))) else error (Errors.msg "Veriloggen: eff_addressing misaligned stack offset") - | _, _ => error (Errors.msg "Veriloggen: translate_arr_access unsuported addressing") + | _, _, _ => error (Errors.msg "Veriloggen: translate_arr_access unsuported addressing") end. Definition transf_instr (fin rtrn stack: reg) (ni: node * instruction) : mon unit := diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index 1f9e368..34fa8ec 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -51,7 +51,7 @@ Inductive match_arrs (m : HTL.module) (f : RTL.function) (sp : Values.val) (mem 0 <= ptr < Z.of_nat m.(HTL.mod_stk_len) -> opt_val_value_lessdef (Mem.loadv AST.Mint32 mem (Values.Val.offset_ptr sp (Integers.Ptrofs.repr ptr))) - (nth (Z.to_nat ptr) stack (ZToValue 32 0))) -> + (nth (Z.to_nat ptr / 4) stack (ZToValue 32 0))) -> match_arrs m f sp mem asa. Inductive match_stacks (mem : mem) : list RTL.stackframe -> list HTL.stackframe -> Prop := @@ -425,38 +425,113 @@ Section CORRECTNESS. | [ H: Values.Vptr _ _ = Values.Vptr _ _ |- _] => invert H end. - - destruct c, addr, args; simplify; rt; t; simplify. + Opaque Nat.div. + + - destruct c, chunk, addr, args; simplify; rt; simplify. + admit. (* FIXME: Alignment *) + + admit. + + admit. + + invert MARR. simplify. - + unfold Op.eval_addressing in *. - (* destruct (Archi.ptr64); simplify; *) - (* destruct (Registers.Regmap.get r0 rs); simplify. *) - admit. + unfold Op.eval_addressing in H0. + destruct (Archi.ptr64) eqn:ARCHI; simplify. + rewrite ARCHI in H0. simplify. - (* eexists. split. *) - (* eapply Smallstep.plus_one. *) - (* eapply HTL.step_module; eauto. *) - (* econstructor. econstructor. econstructor. simpl. *) - (* econstructor. econstructor. econstructor. econstructor. econstructor. *) - (* econstructor. econstructor. econstructor. econstructor. econstructor. *) - (* econstructor. econstructor. econstructor. simpl. *) + (* Out of bounds reads are undefined behaviour. *) + assert (forall ptr v, f.(RTL.fn_stacksize) <= ptr -> + Mem.loadv AST.Mint32 m (Values.Val.offset_ptr sp (Integers.Ptrofs.repr ptr)) = Some v -> + v = Values.Vundef) as INVALID_READ by admit. - (* inversion MARR; subst; clear MARR. simpl in *. *) - (* unfold Verilog.arr_assocmap_lookup. *) - (* rewrite H3. *) + (* Case split on whether read is out of bounds. *) + destruct (RTL.fn_stacksize f <=? Integers.Ptrofs.unsigned i0) eqn:BOUND. + * assert (RTL.fn_stacksize f <= Integers.Ptrofs.unsigned i0) as OUT_OF_BOUNDS. { + apply Zle_bool_imp_le. assumption. + } + eapply INVALID_READ in OUT_OF_BOUNDS. + 2: { rewrite Integers.Ptrofs.repr_unsigned. eassumption. } - (* admit. *) + eexists. split. + eapply Smallstep.plus_one. + eapply HTL.step_module; eauto. + econstructor. econstructor. econstructor. simplify. + econstructor. econstructor. econstructor. econstructor. simplify. - + unfold Op.eval_addressing in *. - (* destruct (Archi.ptr64); simplify; *) - (* destruct (Registers.Regmap.get r0 rs); *) - (* destruct (Registers.Regmap.get r1 rs); simplify; *) - (* destruct (Archi.ptr64); simplify. *) - admit. + unfold Verilog.arr_assocmap_lookup. rewrite H3. + reflexivity. - + admit. + simplify. unfold_merge. apply AssocMap.gss. + + simplify. rewrite assumption_32bit. + constructor. + + unfold_merge. + apply regs_lessdef_add_greater. + apply greater_than_max_func. + + apply regs_lessdef_add_match. + rewrite OUT_OF_BOUNDS. + constructor. assumption. + + assumption. + + unfold state_st_wf. inversion 1. simplify. + unfold_merge. apply AssocMap.gss. + + assumption. + + econstructor; simplify; try reflexivity; eassumption. + + * assert (Integers.Ptrofs.unsigned i0 < f.(RTL.fn_stacksize)) as IN_BOUNDS. { + rewrite Z.leb_antisym in BOUND. + rewrite negb_false_iff in BOUND. + apply Zlt_is_lt_bool. assumption. + } + + eexists. split. + eapply Smallstep.plus_one. + eapply HTL.step_module; eauto. + econstructor. econstructor. econstructor. simplify. + econstructor. econstructor. econstructor. econstructor. simplify. + + unfold Verilog.arr_assocmap_lookup. rewrite H3. + reflexivity. + + simplify. unfold_merge. apply AssocMap.gss. + + simplify. rewrite assumption_32bit. + constructor. + + unfold_merge. + apply regs_lessdef_add_greater. + apply greater_than_max_func. + + apply regs_lessdef_add_match. + + assert (exists ptr, (Z.to_nat ptr / 4)%nat = (valueToNat (ZToValue 32 (Integers.Ptrofs.unsigned i0 / 4)))) + by admit. + simplify. rewrite <- H5. + specialize (H4 x). + assert (0 <= x < Z.of_nat (Z.to_nat (RTL.fn_stacksize f / 4))) by admit. + apply H4 in H0. + invert H0. + assert (Integers.Ptrofs.repr x = i0) by admit. + rewrite H0 in H6. + rewrite H1 in H6. + invert H6. + assumption. + + assumption. + + assumption. + + unfold state_st_wf. inversion 1. simplify. + unfold_merge. apply AssocMap.gss. + + assumption. + + econstructor; simplify; try reflexivity; eassumption. - admit. diff --git a/src/verilog/Verilog.v b/src/verilog/Verilog.v index 845d706..0e999de 100644 --- a/src/verilog/Verilog.v +++ b/src/verilog/Verilog.v @@ -56,10 +56,10 @@ Definition merge_associations {A : Type} (assoc : associations A) := mkassociations (AssocMapExt.merge A assoc.(assoc_nonblocking) assoc.(assoc_blocking)) (AssocMap.empty A). -Definition arr_assocmap_lookup (a : assocmap_arr) (r : reg) (i : nat) : option value := +Definition arr_assocmap_lookup (a : assocmap_arr) (r : reg) (i : nat) : value := match a ! r with - | None => None - | Some arr => nth_error arr i + | None => natToValue 32 0 + | Some arr => nth i arr (natToValue 32 0) end. Fixpoint list_set {A : Type} (i : nat) (x : A) (l : list A) : list A := @@ -297,7 +297,7 @@ Inductive expr_runp : fext -> assocmap -> assocmap_arr -> expr -> value -> Prop | erun_Vvari : forall fext reg stack v iexp i r, expr_runp fext reg stack iexp i -> - arr_assocmap_lookup stack r (valueToNat i) = Some v -> + arr_assocmap_lookup stack r (valueToNat i) = v -> expr_runp fext reg stack (Vvari r iexp) v | erun_Vinputvar : forall fext reg stack r v, -- cgit From 9acb804500b590edbff66cd802216f58dde169cd Mon Sep 17 00:00:00 2001 From: James Pollard Date: Fri, 12 Jun 2020 16:54:28 +0100 Subject: Some work on store proof. --- src/translation/HTLgenproof.v | 49 +++++++++++++++++++++++++++++++++---------- 1 file changed, 38 insertions(+), 11 deletions(-) diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index 34fa8ec..6dc4c21 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -415,16 +415,6 @@ Section CORRECTNESS. | [ _ : context[match ?x with | _ => _ end] |- _ ] => destruct x end. - Ltac t := - match goal with - | [ _ : Mem.loadv _ _ ?a = Some _ |- _ ] => - let PTR := fresh "PTR" in - assert (exists b ofs, a = Values.Vptr b ofs) as PTR; - [> destruct a; simpl in *; try discriminate; - repeat eexists |] - | [ H: Values.Vptr _ _ = Values.Vptr _ _ |- _] => invert H - end. - Opaque Nat.div. - destruct c, chunk, addr, args; simplify; rt; simplify. @@ -522,6 +512,11 @@ Section CORRECTNESS. invert H6. assumption. + assert (Integers.Ptrofs.repr x = i0) by admit. + rewrite H0 in H7. + rewrite H1 in H7. + discriminate. + assumption. assumption. @@ -533,7 +528,39 @@ Section CORRECTNESS. econstructor; simplify; try reflexivity; eassumption. - - admit. + - destruct c, chunk, addr, args; simplify; rt; simplify. + + admit. + + admit. + + admit. + + + eexists. split. + eapply Smallstep.plus_one. + eapply HTL.step_module; eauto. + econstructor. econstructor. econstructor. simplify. + eapply Verilog.stmnt_runp_Vblock_arr. simplify. + econstructor. econstructor. econstructor. simplify. + + reflexivity. + + unfold_merge. apply AssocMap.gss. + + simplify. rewrite assumption_32bit. econstructor. + + unfold_merge. + apply regs_lessdef_add_greater. + apply greater_than_max_func. + assumption. + + assumption. + + unfold state_st_wf. inversion 1. simplify. + unfold_merge. apply AssocMap.gss. + + admit. + + econstructor; simplify; try reflexivity. + admit. + - eexists. split. apply Smallstep.plus_one. eapply HTL.step_module; eauto. -- cgit From eacdec2dd13611f94fe12a41cf04cf38dc389092 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Fri, 12 Jun 2020 18:13:52 +0100 Subject: Fix broken proof. --- src/translation/HTLgenspec.v | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/src/translation/HTLgenspec.v b/src/translation/HTLgenspec.v index 57d2d87..57d7d62 100644 --- a/src/translation/HTLgenspec.v +++ b/src/translation/HTLgenspec.v @@ -433,8 +433,8 @@ Proof. Qed. Hint Resolve iter_expand_instr_spec : htlspec. -Lemma create_arr_inv : forall x y z a b c d, - create_arr x y z = OK (a, b) c d -> y = b. +Lemma create_arr_inv : forall w x y z a b c d, + create_arr w x y z = OK (a, b) c d -> y = b. Proof. inversion 1. reflexivity. Qed. @@ -454,14 +454,14 @@ Proof. monadInv Heqr. (* TODO: We should be able to fold this into the automation. *) - pose proof (create_arr_inv _ _ _ _ _ _ _ EQ0) as STK_LEN. + pose proof (create_arr_inv _ _ _ _ _ _ _ _ EQ0) as STK_LEN. rewrite <- STK_LEN. econstructor; simpl; trivial. intros. inv_incr. assert (EQ3D := EQ3). - destruct x3. + destruct x4. apply collect_declare_datapath_trans in EQ3. apply collect_declare_controllogic_trans in EQ3D. assert (STC: st_controllogic s10 = st_controllogic s3) by congruence. -- cgit From 21ae45cd03f7d38b1ef12270307274a9ee370e17 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Sat, 13 Jun 2020 21:36:32 +0100 Subject: Add a basic length-indexed list type. --- src/verilog/Array.v | 123 ++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 123 insertions(+) create mode 100644 src/verilog/Array.v diff --git a/src/verilog/Array.v b/src/verilog/Array.v new file mode 100644 index 0000000..8673f0c --- /dev/null +++ b/src/verilog/Array.v @@ -0,0 +1,123 @@ +Set Implicit Arguments. + +Require Import Lia. +Require Import Coquplib. +From Coq Require Import Lists.List Datatypes. + +Import ListNotations. + +Local Open Scope nat_scope. + +Record Array (A : Type) : Type := + mk_array + { arr_contents : list A + ; arr_length : nat + ; arr_wf : length arr_contents = arr_length + }. + +Definition make_array {A : Type} (l : list A) : Array A := + @mk_array A l (length l) eq_refl. + +Fixpoint list_set {A : Type} (i : nat) (x : A) (l : list A) {struct l} : list A := + match i, l with + | _, nil => nil + | S n, h :: t => h :: list_set n x t + | O, h :: t => x :: t + end. + +Lemma list_set_spec1 {A : Type} : + forall l i (x : A), + i < length l -> nth_error (list_set i x l) i = Some x. +Proof. + induction l; intros; destruct i; simplify; try lia; try reflexivity; firstorder. +Qed. +Hint Resolve list_set_spec1 : array. + +Lemma list_set_spec2 {A : Type} : + forall l i (x : A) d, + i < length l -> nth i (list_set i x l) d = x. +Proof. + induction l; intros; destruct i; simplify; try lia; try reflexivity; firstorder. +Qed. +Hint Resolve list_set_spec2 : array. + +Lemma array_set_wf {A : Type} : + forall l ln i (x : A), + length l = ln -> length (list_set i x l) = ln. +Proof. + induction l; intros; destruct i; auto. + + invert H; simplify; auto. +Qed. + +Definition array_set {A : Type} (i : nat) (x : A) (a : Array A) := + let l := a.(arr_contents) in + let ln := a.(arr_length) in + let WF := a.(arr_wf) in + @mk_array A (list_set i x l) ln (@array_set_wf A l ln i x WF). + +Lemma array_set_spec1 {A : Type} : + forall a i (x : A), + i < a.(arr_length) -> nth_error ((array_set i x a).(arr_contents)) i = Some x. +Proof. + intros. + + rewrite <- a.(arr_wf) in H. + unfold array_set. simplify. + eauto with array. +Qed. +Hint Resolve array_set_spec1 : array. + +Lemma array_set_spec2 {A : Type} : + forall a i (x : A) d, + i < a.(arr_length) -> nth i ((array_set i x a).(arr_contents)) d = x. +Proof. + intros. + + rewrite <- a.(arr_wf) in H. + unfold array_set. simplify. + eauto with array. +Qed. +Hint Resolve array_set_spec2 : array. + +Definition array_get_error {A : Type} (i : nat) (a : Array A) : option A := + nth_error a.(arr_contents) i. + +Lemma array_get_error_bound {A : Type} : + forall (a : Array A) i, + i < a.(arr_length) -> exists x, array_get_error i a = Some x. +Proof. + intros. + + rewrite <- a.(arr_wf) in H. + assert (~ length (arr_contents a) <= i) by lia. + + pose proof (nth_error_None a.(arr_contents) i). + apply not_iff_compat in H1. + apply <- H1 in H0. + + destruct (nth_error (arr_contents a) i ) eqn:EQ; try contradiction; eauto. +Qed. + +Lemma array_get_error_set_bound {A : Type} : + forall (a : Array A) i x, + i < a.(arr_length) -> array_get_error i (array_set i x a) = Some x. +Proof. + intros. + + unfold array_get_error. + eauto with array. +Qed. + +Definition array_get {A : Type} (i : nat) (x : A) (a : Array A) : A := + nth i a.(arr_contents) x. + +Lemma array_get_set_bound {A : Type} : + forall (a : Array A) i x d, + i < a.(arr_length) -> array_get i d (array_set i x a) = x. +Proof. + intros. + + unfold array_get. + info_eauto with array. +Qed. -- cgit From 39d438f9c2b3d1484ae0e2afe33a19c2f654a8b0 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Sun, 14 Jun 2020 15:27:37 +0100 Subject: Move some standard tactics to Coquplib. --- src/common/Coquplib.v | 10 ++++++++++ 1 file changed, 10 insertions(+) diff --git a/src/common/Coquplib.v b/src/common/Coquplib.v index 47360d6..fd0987b 100644 --- a/src/common/Coquplib.v +++ b/src/common/Coquplib.v @@ -41,6 +41,16 @@ Ltac solve_by_inverts n := Ltac solve_by_invert := solve_by_inverts 1. +Ltac invert x := inversion x; subst; clear x. + +Ltac clear_obvious := + repeat match goal with + | [ H : ex _ |- _ ] => invert H + | [ H : ?C _ = ?C _ |- _ ] => invert H + end. + +Ltac simplify := simpl in *; clear_obvious; simpl in *; try discriminate. + (* Definition const (A B : Type) (a : A) (b : B) : A := a. Definition compose (A B C : Type) (f : B -> C) (g : A -> B) (x : A) : C := f (g x). *) -- cgit From 044a68b1b215125e2651c637f28c794536d27ba5 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Sun, 14 Jun 2020 16:41:27 +0100 Subject: Array semantics now uses dependent Array type. --- _CoqProject | 2 +- src/translation/HTLgen.v | 14 +++++------ src/translation/Veriloggen.v | 5 ++-- src/verilog/HTL.v | 21 +++++++--------- src/verilog/Verilog.v | 57 ++++++++++++++++++++++++-------------------- 5 files changed, 50 insertions(+), 49 deletions(-) diff --git a/_CoqProject b/_CoqProject index a1026ed..7965da9 100644 --- a/_CoqProject +++ b/_CoqProject @@ -14,4 +14,4 @@ -R lib/CompCert/flocq compcert.flocq -R lib/CompCert/lib compcert.lib -R lib/CompCert/x86 compcert.x86 --R lib/CompCert/x86_32 compcert.x86_64 +-R lib/CompCert/x86_32 compcert.x86_32 diff --git a/src/translation/HTLgen.v b/src/translation/HTLgen.v index d35a296..1c67fe7 100644 --- a/src/translation/HTLgen.v +++ b/src/translation/HTLgen.v @@ -131,7 +131,7 @@ Lemma declare_reg_state_incr : s.(st_st) s.(st_freshreg) s.(st_freshstate) - (AssocMap.set r (i, Scalar sz) s.(st_scldecls)) + (AssocMap.set r (i, VScalar sz) s.(st_scldecls)) s.(st_arrdecls) s.(st_datapath) s.(st_controllogic)). @@ -142,7 +142,7 @@ Definition declare_reg (i : option io) (r : reg) (sz : nat) : mon unit := s.(st_st) s.(st_freshreg) s.(st_freshstate) - (AssocMap.set r (i, Scalar sz) s.(st_scldecls)) + (AssocMap.set r (i, VScalar sz) s.(st_scldecls)) s.(st_arrdecls) s.(st_datapath) s.(st_controllogic)) @@ -393,7 +393,7 @@ Lemma create_reg_state_incr: s.(st_st) (Pos.succ (st_freshreg s)) (st_freshstate s) - (AssocMap.set s.(st_freshreg) (i, Scalar sz) s.(st_scldecls)) + (AssocMap.set s.(st_freshreg) (i, VScalar sz) s.(st_scldecls)) s.(st_arrdecls) (st_datapath s) (st_controllogic s)). @@ -405,7 +405,7 @@ Definition create_reg (i : option io) (sz : nat) : mon reg := s.(st_st) (Pos.succ r) (st_freshstate s) - (AssocMap.set s.(st_freshreg) (i, Scalar sz) s.(st_scldecls)) + (AssocMap.set s.(st_freshreg) (i, VScalar sz) s.(st_scldecls)) (st_arrdecls s) (st_datapath s) (st_controllogic s)) @@ -418,7 +418,7 @@ Lemma create_arr_state_incr: (Pos.succ (st_freshreg s)) (st_freshstate s) s.(st_scldecls) - (AssocMap.set s.(st_freshreg) (i, Array sz ln) s.(st_arrdecls)) + (AssocMap.set s.(st_freshreg) (i, VArray sz ln) s.(st_arrdecls)) (st_datapath s) (st_controllogic s)). Proof. constructor; simpl; auto with htlh. Qed. @@ -430,7 +430,7 @@ Definition create_arr (i : option io) (sz : nat) (ln : nat) : mon (reg * nat) := (Pos.succ r) (st_freshstate s) s.(st_scldecls) - (AssocMap.set s.(st_freshreg) (i, Array sz ln) s.(st_arrdecls)) + (AssocMap.set s.(st_freshreg) (i, VArray sz ln) s.(st_arrdecls)) (st_datapath s) (st_controllogic s)) (create_arr_state_incr s sz ln i). @@ -466,7 +466,7 @@ Definition max_state (f: function) : state := mkstate st (Pos.succ st) (Pos.succ (max_pc_function f)) - (AssocMap.set st (None, Scalar 32) (st_scldecls (init_state st))) + (AssocMap.set st (None, VScalar 32) (st_scldecls (init_state st))) (st_arrdecls (init_state st)) (st_datapath (init_state st)) (st_controllogic (init_state st)). diff --git a/src/translation/Veriloggen.v b/src/translation/Veriloggen.v index efe3565..2b9974b 100644 --- a/src/translation/Veriloggen.v +++ b/src/translation/Veriloggen.v @@ -29,13 +29,13 @@ Fixpoint transl_list (st : list (node * Verilog.stmnt)) {struct st} : list (expr Fixpoint scl_to_Vdecl (scldecl : list (reg * (option io * scl_decl))) {struct scldecl} : list module_item := match scldecl with - | (r, (io, Scalar sz))::scldecl' => Vdeclaration (Vdecl io r sz) :: scl_to_Vdecl scldecl' + | (r, (io, VScalar sz))::scldecl' => Vdeclaration (Vdecl io r sz) :: scl_to_Vdecl scldecl' | nil => nil end. Fixpoint arr_to_Vdeclarr (arrdecl : list (reg * (option io * arr_decl))) {struct arrdecl} : list module_item := match arrdecl with - | (r, (io, Array sz l))::arrdecl' => Vdeclaration (Vdeclarr io r sz l) :: arr_to_Vdeclarr arrdecl' + | (r, (io, VArray sz l))::arrdecl' => Vdeclaration (Vdeclarr io r sz l) :: arr_to_Vdeclarr arrdecl' | nil => nil end. @@ -56,6 +56,7 @@ Definition transl_module (m : HTL.module) : Verilog.module := m.(mod_return) m.(mod_st) m.(mod_stk) + m.(mod_stk_len) m.(mod_params) body m.(mod_entrypoint). diff --git a/src/verilog/HTL.v b/src/verilog/HTL.v index c509248..8149ddf 100644 --- a/src/verilog/HTL.v +++ b/src/verilog/HTL.v @@ -17,7 +17,7 @@ *) From Coq Require Import FSets.FMapPositive. -From coqup Require Import Coquplib Value AssocMap. +From coqup Require Import Coquplib Value AssocMap Array. From coqup Require Verilog. From compcert Require Events Globalenvs Smallstep Integers Values. From compcert Require Import Maps. @@ -66,13 +66,8 @@ Fixpoint init_regs (vl : list value) (rl : list reg) {struct rl} := | _, _ => empty_assocmap end. -Fixpoint zeroes' (acc : list value) (n : nat) : list value := - match n with - | O => acc - | S n => zeroes' ((NToValue 32 0)::acc) n - end. - -Definition zeroes : nat -> list value := zeroes' nil. +Definition empty_stack (m : module) : AssocMap.t (Array value) := + (AssocMap.set m.(mod_stk) (Verilog.zeroes m.(mod_stk_len)) (AssocMap.empty (Array value))). (** * Operational Semantics *) @@ -84,7 +79,7 @@ Inductive stackframe : Type := (m : module) (pc : node) (reg_assoc : assocmap) - (arr_assoc : AssocMap.t (list value)), + (arr_assoc : AssocMap.t (Array value)), stackframe. Inductive state : Type := @@ -93,7 +88,7 @@ Inductive state : Type := (m : module) (st : node) (reg_assoc : assocmap) - (arr_assoc : AssocMap.t (list value)), state + (arr_assoc : AssocMap.t (Array value)), state | Returnstate : forall (res : list stackframe) (v : value), state @@ -114,7 +109,7 @@ Inductive step : genv -> state -> Events.trace -> state -> Prop := m.(mod_datapath)!st = Some data -> Verilog.stmnt_runp f (Verilog.mkassociations asr empty_assocmap) - (Verilog.mkassociations asa (AssocMap.empty (list value))) + (Verilog.mkassociations asa (empty_stack m)) ctrl (Verilog.mkassociations basr1 nasr1) (Verilog.mkassociations basa1 nasa1) -> @@ -125,7 +120,7 @@ Inductive step : genv -> state -> Events.trace -> state -> Prop := (Verilog.mkassociations basr2 nasr2) (Verilog.mkassociations basa2 nasa2) -> asr' = merge_assocmap nasr2 basr2 -> - asa' = AssocMapExt.merge (list value) nasa2 basa2 -> + asa' = AssocMapExt.merge (Array value) nasa2 basa2 -> asr'!(m.(mod_st)) = Some stval -> valueToPos stval = pstval -> step g (State sf m st asr asa) Events.E0 (State sf m pstval asr' asa') @@ -140,7 +135,7 @@ Inductive step : genv -> state -> Events.trace -> state -> Prop := (State res m m.(mod_entrypoint) (AssocMap.set (mod_st m) (posToValue 32 m.(mod_entrypoint)) (init_regs args m.(mod_params))) - (AssocMap.set m.(mod_stk) (zeroes m.(mod_stk_len)) (AssocMap.empty (list value)))) + (empty_stack m)) | step_return : forall g m asr asa i r sf pc mst, mst = mod_st m -> diff --git a/src/verilog/Verilog.v b/src/verilog/Verilog.v index 8b83d49..3ab3f10 100644 --- a/src/verilog/Verilog.v +++ b/src/verilog/Verilog.v @@ -27,7 +27,7 @@ From Coq Require Import Import ListNotations. -From coqup Require Import common.Coquplib common.Show verilog.Value AssocMap. +From coqup Require Import common.Coquplib common.Show verilog.Value AssocMap Array. From compcert Require Integers Events. From compcert Require Import Errors Smallstep Globalenvs. @@ -48,38 +48,31 @@ Record associations (A : Type) : Type := }. Definition reg_associations := associations value. -Definition arr_associations := associations (list value). +Definition arr_associations := associations (Array value). -Definition assocmap_arr := AssocMap.t (list value). +Definition assocmap_arr := AssocMap.t (Array value). Definition merge_associations {A : Type} (assoc : associations A) := mkassociations (AssocMapExt.merge A assoc.(assoc_nonblocking) assoc.(assoc_blocking)) (AssocMap.empty A). -Definition arr_assocmap_lookup (a : assocmap_arr) (r : reg) (i : nat) : value := +Definition arr_assocmap_lookup (a : assocmap_arr) (r : reg) (i : nat) : option value := match a ! r with - | None => natToValue 32 0 - | Some arr => nth i arr (natToValue 32 0) + | None => None + | Some arr => array_get_error i arr end. -Fixpoint list_set {A : Type} (i : nat) (x : A) (l : list A) : list A := - match i, l with - | _, nil => nil - | S n, h :: t => h :: list_set n x t - | O, h :: t => x :: t - end. - -Definition assocmap_l_set (r : reg) (i : nat) (v : value) (a : assocmap_arr) : assocmap_arr := +Definition arr_assocmap_set (r : reg) (i : nat) (v : value) (a : assocmap_arr) : assocmap_arr := match a ! r with | None => a - | Some arr => AssocMap.set r (list_set i v arr) a + | Some arr => a # r <- (array_set i v arr) end. Definition block_arr (r : reg) (i : nat) (asa : arr_associations) (v : value) : arr_associations := - mkassociations (assocmap_l_set r i v asa.(assoc_blocking)) asa.(assoc_nonblocking). + mkassociations (arr_assocmap_set r i v asa.(assoc_blocking)) asa.(assoc_nonblocking). Definition nonblock_arr (r : reg) (i : nat) (asa : arr_associations) (v : value) : arr_associations := - mkassociations asa.(assoc_blocking) (assocmap_l_set r i v asa.(assoc_nonblocking)). + mkassociations asa.(assoc_blocking) (arr_assocmap_set r i v asa.(assoc_nonblocking)). Definition block_reg (r : reg) (asr : reg_associations) (v : value) := mkassociations (AssocMap.set r v asr.(assoc_blocking)) asr.(assoc_nonblocking). @@ -87,8 +80,8 @@ Definition block_reg (r : reg) (asr : reg_associations) (v : value) := Definition nonblock_reg (r : reg) (asr : reg_associations) (v : value) := mkassociations asr.(assoc_blocking) (AssocMap.set r v asr.(assoc_nonblocking)). -Inductive scl_decl : Type := Scalar (sz : nat). -Inductive arr_decl : Type := Array (sz : nat) (ln : nat). +Inductive scl_decl : Type := VScalar (sz : nat). +Inductive arr_decl : Type := VArray (sz : nat) (ln : nat). (** * Verilog AST @@ -217,6 +210,7 @@ Record module : Type := mkmodule { mod_return : reg; mod_st : reg; (**r Variable that defines the current state, it should be internal. *) mod_stk : reg; + mod_stk_len : nat; mod_args : list reg; mod_body : list module_item; mod_entrypoint : node; @@ -234,7 +228,7 @@ Definition posToLit (p : positive) : expr := Coercion Vlit : value >-> expr. Coercion Vvar : reg >-> expr. -Definition fext := AssocMap.t value. +Definition fext := assocmap. Definition fextclk := nat -> fext. (** ** State @@ -272,7 +266,7 @@ Inductive state : Type := (m : module) (st : node) (reg_assoc : assocmap) - (arr_assoc : AssocMap.t (list value)), state + (arr_assoc : AssocMap.t (Array value)), state | Returnstate : forall (res : list stackframe) (v : value), state @@ -324,7 +318,7 @@ Inductive expr_runp : fext -> assocmap -> assocmap_arr -> expr -> value -> Prop | erun_Vvari : forall fext reg stack v iexp i r, expr_runp fext reg stack iexp i -> - arr_assocmap_lookup stack r (valueToNat i) = v -> + arr_assocmap_lookup stack r (valueToNat i) = Some v -> expr_runp fext reg stack (Vvari r iexp) v | erun_Vinputvar : forall fext reg stack r v, @@ -690,16 +684,27 @@ Fixpoint init_params (vl : list value) (rl : list reg) {struct rl} := Definition genv := Globalenvs.Genv.t fundef unit. +Fixpoint list_zeroes' (acc : list value) (n : nat) : list value := + match n with + | O => acc + | S n => list_zeroes' ((NToValue 32 0)::acc) n + end. + +Definition list_zeroes : nat -> list value := list_zeroes' nil. +Definition zeroes (n : nat) : Array value := make_array (list_zeroes n). +Definition empty_stack (m : module) : AssocMap.t (Array value) := + (AssocMap.set m.(mod_stk) (zeroes m.(mod_stk_len)) (AssocMap.empty (Array value))). + Inductive step : genv -> state -> Events.trace -> state -> Prop := | step_module : forall asr asa asr' asa' basr1 nasr1 basa1 nasa1 f stval pstval m sf st g, mis_stepp f (mkassociations asr empty_assocmap) - (mkassociations asa (AssocMap.empty (list value))) + (mkassociations asa (empty_stack m)) m.(mod_body) (mkassociations basr1 nasr1) (mkassociations basa1 nasa1)-> asr' = merge_assocmap nasr1 basr1 -> - asa' = AssocMapExt.merge (list value) nasa1 basa1 -> + asa' = AssocMapExt.merge (Array value) nasa1 basa1 -> asr'!(m.(mod_st)) = Some stval -> valueToPos stval = pstval -> step g (State sf m st asr asa) Events.E0 (State sf m pstval asr' asa') @@ -714,13 +719,13 @@ Inductive step : genv -> state -> Events.trace -> state -> Prop := (State res m m.(mod_entrypoint) (AssocMap.set m.(mod_st) (posToValue 32 m.(mod_entrypoint)) (init_params args m.(mod_args))) - (AssocMap.empty (list value))) + (empty_stack m)) | step_return : forall g m asr i r sf pc mst, mst = mod_st m -> step g (Returnstate (Stackframe r m pc asr :: sf) i) Events.E0 (State sf m pc ((asr # mst <- (posToValue 32 pc)) # r <- i) - (AssocMap.empty (list value))). + (empty_stack m)). Hint Constructors step : verilog. Inductive initial_state (p: program): state -> Prop := -- cgit From dfea5f0f6307177a9127ce29db496a819dcdb232 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Wed, 17 Jun 2020 19:53:43 +0100 Subject: Fix array semantics merge granularity. --- src/common/Coquplib.v | 7 +++++++ src/verilog/Array.v | 38 ++++++++++++++++++++++++++++++++++ src/verilog/HTL.v | 16 +++++++-------- src/verilog/Verilog.v | 56 +++++++++++++++++++++++++++++++-------------------- 4 files changed, 87 insertions(+), 30 deletions(-) diff --git a/src/common/Coquplib.v b/src/common/Coquplib.v index fd0987b..675ad23 100644 --- a/src/common/Coquplib.v +++ b/src/common/Coquplib.v @@ -47,6 +47,7 @@ Ltac clear_obvious := repeat match goal with | [ H : ex _ |- _ ] => invert H | [ H : ?C _ = ?C _ |- _ ] => invert H + | [ H : _ /\ _ |- _ ] => invert H end. Ltac simplify := simpl in *; clear_obvious; simpl in *; try discriminate. @@ -81,6 +82,12 @@ Definition bind {A B : Type} (f : option A) (g : A -> option B) : option B := | _ => None end. +Definition join {A : Type} (a : option (option A)) : option A := + match a with + | None => None + | Some a' => a' + end. + Module Notation. Notation "'do' X <- A ; B" := (bind A (fun X => B)) (at level 200, X ident, A at level 100, B at level 200). diff --git a/src/verilog/Array.v b/src/verilog/Array.v index 8673f0c..16f5406 100644 --- a/src/verilog/Array.v +++ b/src/verilog/Array.v @@ -121,3 +121,41 @@ Proof. unfold array_get. info_eauto with array. Qed. + +Fixpoint list_repeat' {A : Type} (acc : list A) (a : A) (n : nat) : list A := + match n with + | O => acc + | S n => list_repeat' (a::acc) a n + end. + +Lemma list_repeat'_len : forall {A : Type} (a : A) n l, + Datatypes.length (list_repeat' l a n) = (n + Datatypes.length l)%nat. +Proof. + induction n; intros; simplify; try reflexivity. + + specialize (IHn (a :: l)). + rewrite IHn. + simplify. + lia. +Qed. + +Definition list_repeat {A : Type} : A -> nat -> list A := list_repeat' nil. + +Lemma list_repeat_len {A : Type} : forall n (a : A), Datatypes.length (list_repeat a n) = n. +Proof. + intros. + unfold list_repeat. + rewrite list_repeat'_len. + simplify. lia. +Qed. + +Definition arr_repeat {A : Type} (a : A) (n : nat) : Array A := make_array (list_repeat a n). + +Fixpoint list_combine {A B C : Type} (f : A -> B -> C) (x : list A) (y : list B) : list C := + match x, y with + | a :: t, b :: t' => f a b :: list_combine f t t' + | _, _ => nil + end. + +Definition combine {A B C : Type} (f : A -> B -> C) (x : Array A) (y : Array B) : Array C := + make_array (list_combine f x.(arr_contents) y.(arr_contents)). diff --git a/src/verilog/HTL.v b/src/verilog/HTL.v index 8149ddf..0bf5072 100644 --- a/src/verilog/HTL.v +++ b/src/verilog/HTL.v @@ -66,8 +66,8 @@ Fixpoint init_regs (vl : list value) (rl : list reg) {struct rl} := | _, _ => empty_assocmap end. -Definition empty_stack (m : module) : AssocMap.t (Array value) := - (AssocMap.set m.(mod_stk) (Verilog.zeroes m.(mod_stk_len)) (AssocMap.empty (Array value))). +Definition empty_stack (m : module) : Verilog.assocmap_arr := + (AssocMap.set m.(mod_stk) (Array.arr_repeat None m.(mod_stk_len)) (AssocMap.empty Verilog.arr)). (** * Operational Semantics *) @@ -78,8 +78,8 @@ Inductive stackframe : Type := forall (res : reg) (m : module) (pc : node) - (reg_assoc : assocmap) - (arr_assoc : AssocMap.t (Array value)), + (reg_assoc : Verilog.assocmap_reg) + (arr_assoc : Verilog.assocmap_arr), stackframe. Inductive state : Type := @@ -87,8 +87,8 @@ Inductive state : Type := forall (stack : list stackframe) (m : module) (st : node) - (reg_assoc : assocmap) - (arr_assoc : AssocMap.t (Array value)), state + (reg_assoc : Verilog.assocmap_reg) + (arr_assoc : Verilog.assocmap_arr), state | Returnstate : forall (res : list stackframe) (v : value), state @@ -119,8 +119,8 @@ Inductive step : genv -> state -> Events.trace -> state -> Prop := data (Verilog.mkassociations basr2 nasr2) (Verilog.mkassociations basa2 nasa2) -> - asr' = merge_assocmap nasr2 basr2 -> - asa' = AssocMapExt.merge (Array value) nasa2 basa2 -> + asr' = Verilog.merge_regs nasr2 basr2 -> + asa' = Verilog.merge_arrs nasa2 basa2 -> asr'!(m.(mod_st)) = Some stval -> valueToPos stval = pstval -> step g (State sf m st asr asa) Events.E0 (State sf m pstval asr' asa') diff --git a/src/verilog/Verilog.v b/src/verilog/Verilog.v index 3ab3f10..d20ffcd 100644 --- a/src/verilog/Verilog.v +++ b/src/verilog/Verilog.v @@ -25,6 +25,8 @@ From Coq Require Import Lists.List Program. +Require Import Lia. + Import ListNotations. From coqup Require Import common.Coquplib common.Show verilog.Value AssocMap Array. @@ -47,25 +49,44 @@ Record associations (A : Type) : Type := assoc_nonblocking : AssocMap.t A }. +Definition arr := (Array (option value)). + Definition reg_associations := associations value. -Definition arr_associations := associations (Array value). +Definition arr_associations := associations arr. + +Definition assocmap_reg := AssocMap.t value. +Definition assocmap_arr := AssocMap.t arr. -Definition assocmap_arr := AssocMap.t (Array value). +Definition merge_regs (new : assocmap_reg) (old : assocmap_reg) : assocmap_reg := + AssocMapExt.merge value new old. -Definition merge_associations {A : Type} (assoc : associations A) := - mkassociations (AssocMapExt.merge A assoc.(assoc_nonblocking) assoc.(assoc_blocking)) - (AssocMap.empty A). +Definition merge_cell (new : option value) (old : option value) : option value := + match new, old with + | Some _, _ => new + | _, _ => old + end. + +Definition merge_arr (new : option arr) (old : option arr) : option arr := + match new, old with + | Some new', Some old' => Some (combine merge_cell new' old') + | Some new', None => Some new' + | None, Some old' => Some old' + | None, None => None + end. + +Definition merge_arrs (new : assocmap_arr) (old : assocmap_arr) : assocmap_arr := + AssocMap.combine merge_arr new old. Definition arr_assocmap_lookup (a : assocmap_arr) (r : reg) (i : nat) : option value := match a ! r with | None => None - | Some arr => array_get_error i arr + | Some arr => Option.join (array_get_error i arr) end. Definition arr_assocmap_set (r : reg) (i : nat) (v : value) (a : assocmap_arr) : assocmap_arr := match a ! r with | None => a - | Some arr => a # r <- (array_set i v arr) + | Some arr => a # r <- (array_set i (Some v) arr) end. Definition block_arr (r : reg) (i : nat) (asa : arr_associations) (v : value) : arr_associations := @@ -265,8 +286,8 @@ Inductive state : Type := forall (stack : list stackframe) (m : module) (st : node) - (reg_assoc : assocmap) - (arr_assoc : AssocMap.t (Array value)), state + (reg_assoc : assocmap_reg) + (arr_assoc : assocmap_arr), state | Returnstate : forall (res : list stackframe) (v : value), state @@ -683,17 +704,8 @@ Fixpoint init_params (vl : list value) (rl : list reg) {struct rl} := end. Definition genv := Globalenvs.Genv.t fundef unit. - -Fixpoint list_zeroes' (acc : list value) (n : nat) : list value := - match n with - | O => acc - | S n => list_zeroes' ((NToValue 32 0)::acc) n - end. - -Definition list_zeroes : nat -> list value := list_zeroes' nil. -Definition zeroes (n : nat) : Array value := make_array (list_zeroes n). -Definition empty_stack (m : module) : AssocMap.t (Array value) := - (AssocMap.set m.(mod_stk) (zeroes m.(mod_stk_len)) (AssocMap.empty (Array value))). +Definition empty_stack (m : module) : assocmap_arr := + (AssocMap.set m.(mod_stk) (Array.arr_repeat None m.(mod_stk_len)) (AssocMap.empty arr)). Inductive step : genv -> state -> Events.trace -> state -> Prop := | step_module : @@ -703,8 +715,8 @@ Inductive step : genv -> state -> Events.trace -> state -> Prop := m.(mod_body) (mkassociations basr1 nasr1) (mkassociations basa1 nasa1)-> - asr' = merge_assocmap nasr1 basr1 -> - asa' = AssocMapExt.merge (Array value) nasa1 basa1 -> + asr' = merge_regs nasr1 basr1 -> + asa' = merge_arrs nasa1 basa1 -> asr'!(m.(mod_st)) = Some stval -> valueToPos stval = pstval -> step g (State sf m st asr asa) Events.E0 (State sf m pstval asr' asa') -- cgit From 58f0022a8b5f9ab42e1a8515a77820a9d086ba76 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Wed, 17 Jun 2020 20:05:05 +0100 Subject: Use NBAs for loads and stores. --- src/translation/HTLgen.v | 4 ++-- src/translation/HTLgenspec.v | 4 ++-- 2 files changed, 4 insertions(+), 4 deletions(-) diff --git a/src/translation/HTLgen.v b/src/translation/HTLgen.v index 1c67fe7..73f2b63 100644 --- a/src/translation/HTLgen.v +++ b/src/translation/HTLgen.v @@ -366,10 +366,10 @@ Definition transf_instr (fin rtrn stack: reg) (ni: node * instruction) : mon uni | Iload mem addr args dst n' => do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; - add_instr n n' (block dst src) + add_instr n n' (nonblock dst src) | Istore mem addr args src n' => do dst <- translate_arr_access mem addr args stack; - add_instr n n' (Vblock dst (Vvar src)) (* TODO: Could juse use add_instr? reg exists. *) + add_instr n n' (Vnonblock dst (Vvar src)) (* TODO: Could juse use add_instr? reg exists. *) | Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") diff --git a/src/translation/HTLgenspec.v b/src/translation/HTLgenspec.v index 57d7d62..1a9144c 100644 --- a/src/translation/HTLgenspec.v +++ b/src/translation/HTLgenspec.v @@ -163,11 +163,11 @@ Inductive tr_instr (fin rtrn st stk : reg) : RTL.instruction -> stmnt -> stmnt - | tr_instr_Iload : forall mem addr args s s' i c dst n, translate_arr_access mem addr args stk s = OK c s' i -> - tr_instr fin rtrn st stk (RTL.Iload mem addr args dst n) (block dst c) (state_goto st n) + tr_instr fin rtrn st stk (RTL.Iload mem addr args dst n) (nonblock dst c) (state_goto st n) | tr_instr_Istore : forall mem addr args s s' i c src n, translate_arr_access mem addr args stk s = OK c s' i -> - tr_instr fin rtrn st stk (RTL.Istore mem addr args src n) (Vblock c (Vvar src)) + tr_instr fin rtrn st stk (RTL.Istore mem addr args src n) (Vnonblock c (Vvar src)) (state_goto st n). Hint Constructors tr_instr : htlspec. -- cgit From 00c579e603478d452959dde0ec61672d7b5d27a4 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Wed, 17 Jun 2020 23:08:32 +0100 Subject: Some (very) useful lemmas about arrays. --- src/common/Coquplib.v | 2 + src/translation/HTLgenproof.v | 36 ++++++++++++-- src/verilog/Array.v | 108 ++++++++++++++++++++++++++++++++++++++++-- 3 files changed, 139 insertions(+), 7 deletions(-) diff --git a/src/common/Coquplib.v b/src/common/Coquplib.v index 675ad23..efa1323 100644 --- a/src/common/Coquplib.v +++ b/src/common/Coquplib.v @@ -52,6 +52,8 @@ Ltac clear_obvious := Ltac simplify := simpl in *; clear_obvious; simpl in *; try discriminate. +Global Opaque Nat.div. + (* Definition const (A B : Type) (a : A) (b : B) : A := a. Definition compose (A B C : Type) (f : B -> C) (g : A -> B) (x : A) : C := f (g x). *) diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index 773497b..1356f08 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -43,15 +43,16 @@ Definition state_st_wf (m : HTL.module) (s : HTL.state) := Hint Unfold state_st_wf : htlproof. Inductive match_arrs (m : HTL.module) (f : RTL.function) (sp : Values.val) (mem : mem) : - AssocMap.t (list value) -> Prop := + Verilog.assocmap_arr -> Prop := | match_arr : forall asa stack, - m.(HTL.mod_stk_len) = Z.to_nat (f.(RTL.fn_stacksize) / 4) -> - asa ! (m.(HTL.mod_stk)) = Some stack -> + asa ! (m.(HTL.mod_stk)) = Some stack /\ + stack.(arr_length) = Z.to_nat (f.(RTL.fn_stacksize) / 4) /\ (forall ptr, 0 <= ptr < Z.of_nat m.(HTL.mod_stk_len) -> opt_val_value_lessdef (Mem.loadv AST.Mint32 mem (Values.Val.offset_ptr sp (Integers.Ptrofs.repr ptr))) - (nth (Z.to_nat ptr / 4) stack (ZToValue 32 0))) -> + (Option.default (NToValue 32 0) + (Option.join (array_get_error (Z.to_nat ptr / 4) stack)))) -> match_arrs m f sp mem asa. Inductive match_stacks (mem : mem) : list RTL.stackframe -> list HTL.stackframe -> Prop := @@ -130,6 +131,33 @@ Proof. Qed. Hint Resolve regs_lessdef_add_match : htlproof. +Lemma list_combine_none : + forall n l, + length l = n -> + list_combine Verilog.merge_cell (list_repeat None n) l = l. +Proof. + induction n; intros; simplify. + - symmetry. apply length_zero_iff_nil. auto. + - destruct l; simplify. + rewrite list_repeat_cons. + simplify. f_equal. + eauto. +Qed. + +Lemma combine_none : + forall n a, + a.(arr_length) = n -> + arr_contents (combine Verilog.merge_cell (arr_repeat None n) a) = arr_contents a. +Proof. + intros. + unfold combine. + simplify. + + rewrite <- (arr_wf a) in H. + apply list_combine_none. + assumption. +Qed. + (* Need to eventually move posToValue 32 to posToValueAuto, as that has this proof. *) Lemma assumption_32bit : forall v, diff --git a/src/verilog/Array.v b/src/verilog/Array.v index 16f5406..0b6e2a9 100644 --- a/src/verilog/Array.v +++ b/src/verilog/Array.v @@ -122,14 +122,15 @@ Proof. info_eauto with array. Qed. +(** Tail recursive version of standard library function. *) Fixpoint list_repeat' {A : Type} (acc : list A) (a : A) (n : nat) : list A := match n with | O => acc | S n => list_repeat' (a::acc) a n end. -Lemma list_repeat'_len : forall {A : Type} (a : A) n l, - Datatypes.length (list_repeat' l a n) = (n + Datatypes.length l)%nat. +Lemma list_repeat'_len {A : Type} : forall (a : A) n l, + length (list_repeat' l a n) = (n + Datatypes.length l)%nat. Proof. induction n; intros; simplify; try reflexivity. @@ -139,9 +140,46 @@ Proof. lia. Qed. +Lemma list_repeat'_app {A : Type} : forall (a : A) n l, + list_repeat' l a n = list_repeat' [] a n ++ l. +Proof. + induction n; intros; simplify; try reflexivity. + + pose proof IHn. + specialize (H (a :: l)). + rewrite H. clear H. + specialize (IHn (a :: nil)). + rewrite IHn. clear IHn. + remember (list_repeat' [] a n) as l0. + + rewrite <- app_assoc. + f_equal. +Qed. + +Lemma list_repeat'_head_tail {A : Type} : forall n (a : A), + a :: list_repeat' [] a n = list_repeat' [] a n ++ [a]. +Proof. + induction n; intros; simplify; try reflexivity. + rewrite list_repeat'_app. + + replace (a :: list_repeat' [] a n ++ [a]) with (list_repeat' [] a n ++ [a] ++ [a]). + 2: { rewrite app_comm_cons. rewrite IHn; auto. + rewrite app_assoc. reflexivity. } + rewrite app_assoc. reflexivity. +Qed. + +Lemma list_repeat'_cons {A : Type} : forall (a : A) n, + list_repeat' [a] a n = a :: list_repeat' [] a n. +Proof. + intros. + + rewrite list_repeat'_head_tail; auto. + apply list_repeat'_app. +Qed. + Definition list_repeat {A : Type} : A -> nat -> list A := list_repeat' nil. -Lemma list_repeat_len {A : Type} : forall n (a : A), Datatypes.length (list_repeat a n) = n. +Lemma list_repeat_len {A : Type} : forall n (a : A), length (list_repeat a n) = n. Proof. intros. unfold list_repeat. @@ -149,6 +187,42 @@ Proof. simplify. lia. Qed. +Lemma dec_list_repeat_spec {A : Type} : forall n (a : A) a', + (forall x x' : A, {x' = x} + {~ x' = x}) -> + In a' (list_repeat a n) -> a' = a. +Proof. + induction n; intros; simplify; try contradiction. + + unfold list_repeat in *. + simplify. + + rewrite list_repeat'_app in H. + pose proof (X a a'). + destruct H0; auto. + + (* This is actually a degenerate case, not an unprovable goal. *) + pose proof (in_app_or (list_repeat' [] a n) ([a])). + apply H0 in H. invert H. + + - eapply IHn in X; eassumption. + - invert H1; contradiction. +Qed. + +Lemma list_repeat_head_tail {A : Type} : forall n (a : A), + a :: list_repeat a n = list_repeat a n ++ [a]. +Proof. + unfold list_repeat. apply list_repeat'_head_tail. +Qed. + +Lemma list_repeat_cons {A : Type} : forall n (a : A), + list_repeat a (S n) = a :: list_repeat a n. +Proof. + intros. + + unfold list_repeat. + apply list_repeat'_cons. +Qed. + Definition arr_repeat {A : Type} (a : A) (n : nat) : Array A := make_array (list_repeat a n). Fixpoint list_combine {A B C : Type} (f : A -> B -> C) (x : list A) (y : list B) : list C := @@ -157,5 +231,33 @@ Fixpoint list_combine {A B C : Type} (f : A -> B -> C) (x : list A) (y : list B) | _, _ => nil end. +Lemma list_combine_length {A B C : Type} (f : A -> B -> C) : forall (x : list A) (y : list B), + length (list_combine f x y) = min (length x) (length y). +Proof. + induction x; intros; simplify; try reflexivity. + + destruct y; simplify; auto. +Qed. + Definition combine {A B C : Type} (f : A -> B -> C) (x : Array A) (y : Array B) : Array C := make_array (list_combine f x.(arr_contents) y.(arr_contents)). + +Lemma combine_length {A B C: Type} : forall x y (f : A -> B -> C), + x.(arr_length) = y.(arr_length) -> arr_length (combine f x y) = x.(arr_length). +Proof. + intros. + + unfold combine. + unfold make_array. + simplify. + + rewrite <- (arr_wf x) in *. + rewrite <- (arr_wf y) in *. + + destruct (arr_contents x); simplify. + - reflexivity. + - destruct (arr_contents y); simplify. + f_equal. + rewrite list_combine_length. + destruct (Min.min_dec (length l) (length l0)); congruence. +Qed. -- cgit From 9a65ca2731adf234f5ce946503314267ced62a44 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Wed, 17 Jun 2020 23:19:48 +0100 Subject: Fix Inop proof to work with new array semantics. --- src/translation/HTLgenproof.v | 43 ++++++++++++++++++++++++++++++++++++------- src/verilog/Array.v | 8 ++++++++ 2 files changed, 44 insertions(+), 7 deletions(-) diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index 1356f08..80d936f 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -18,7 +18,7 @@ From compcert Require RTL Registers AST Integers. From compcert Require Import Globalenvs Memory. -From coqup Require Import Coquplib HTLgenspec HTLgen Value AssocMap. +From coqup Require Import Coquplib HTLgenspec HTLgen Value AssocMap Array. From coqup Require HTL Verilog. Local Open Scope assocmap. @@ -326,21 +326,51 @@ Section CORRECTNESS. eapply HTL.step_module; eauto. (* processing of state *) econstructor. - simpl. trivial. + simplify. econstructor. econstructor. econstructor. + simplify. - (* prove stval *) - unfold merge_assocmap. - unfold_merge. simpl. apply AssocMap.gss. + unfold Verilog.merge_regs. + unfold_merge. apply AssocMap.gss. (* prove match_state *) rewrite assumption_32bit. - constructor; auto. + constructor; auto; simplify. + + unfold Verilog.merge_regs. unfold_merge. simpl. apply regs_lessdef_add_greater. apply greater_than_max_func. assumption. + unfold Verilog.merge_regs. unfold state_st_wf. inversion 1. subst. unfold_merge. apply AssocMap.gss. + + (* prove match_arrs *) + invert MARR. simplify. + unfold HTL.empty_stack. simplify. unfold Verilog.merge_arrs. + econstructor. + simplify. repeat split. + + rewrite AssocMap.gcombine. + 2: { reflexivity. } + rewrite AssocMap.gss. + unfold Verilog.merge_arr. + setoid_rewrite H1. + reflexivity. + + rewrite combine_length. + unfold arr_repeat. simplify. + rewrite list_repeat_len. + reflexivity. + + unfold arr_repeat. simplify. + rewrite list_repeat_len; auto. + intros. + erewrite array_get_error_equal. + eauto. apply combine_none. + + assumption. + - (* Iop *) (* destruct v eqn:?; *) (* try ( *) @@ -443,7 +473,6 @@ Section CORRECTNESS. | [ _ : context[match ?x with | _ => _ end] |- _ ] => destruct x end. - Opaque Nat.div. - destruct c, chunk, addr, args; simplify; rt; simplify. diff --git a/src/verilog/Array.v b/src/verilog/Array.v index 0b6e2a9..705dea7 100644 --- a/src/verilog/Array.v +++ b/src/verilog/Array.v @@ -83,6 +83,14 @@ Hint Resolve array_set_spec2 : array. Definition array_get_error {A : Type} (i : nat) (a : Array A) : option A := nth_error a.(arr_contents) i. +Lemma array_get_error_equal {A : Type} : + forall (a b : Array A) i, + a.(arr_contents) = b.(arr_contents) -> + array_get_error i a = array_get_error i b. +Proof. + unfold array_get_error. congruence. +Qed. + Lemma array_get_error_bound {A : Type} : forall (a : Array A) i, i < a.(arr_length) -> exists x, array_get_error i a = Some x. -- cgit From 9e49a65aa01e79b85a35d1dd15f45ee89e3e9906 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Thu, 18 Jun 2020 00:01:11 +0100 Subject: Fix array semantics behaviour for undefined values. --- src/verilog/Verilog.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/verilog/Verilog.v b/src/verilog/Verilog.v index d20ffcd..4144632 100644 --- a/src/verilog/Verilog.v +++ b/src/verilog/Verilog.v @@ -80,7 +80,7 @@ Definition merge_arrs (new : assocmap_arr) (old : assocmap_arr) : assocmap_arr : Definition arr_assocmap_lookup (a : assocmap_arr) (r : reg) (i : nat) : option value := match a ! r with | None => None - | Some arr => Option.join (array_get_error i arr) + | Some arr => Some (Option.default (NToValue 32 0) (Option.join (array_get_error i arr))) end. Definition arr_assocmap_set (r : reg) (i : nat) (v : value) (a : assocmap_arr) : assocmap_arr := -- cgit From 39c336a3e507b9264cd80d1721b724dc5606de6d Mon Sep 17 00:00:00 2001 From: James Pollard Date: Thu, 18 Jun 2020 00:47:29 +0100 Subject: Fix up ILoad proof. --- src/translation/HTLgenproof.v | 153 ++++++++++++++++-------------------------- 1 file changed, 58 insertions(+), 95 deletions(-) diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index 80d936f..8a246f3 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -21,6 +21,8 @@ From compcert Require Import Globalenvs Memory. From coqup Require Import Coquplib HTLgenspec HTLgen Value AssocMap Array. From coqup Require HTL Verilog. +Require Import Lia. + Local Open Scope assocmap. Hint Resolve Smallstep.forward_simulation_plus : htlproof. @@ -449,16 +451,6 @@ Section CORRECTNESS. (* assumption. *) admit. - Ltac invert x := inversion x; subst; clear x. - - Ltac clear_obvious := - repeat match goal with - | [ H : ex _ |- _ ] => invert H - | [ H : Some _ = Some _ |- _ ] => invert H - end. - - Ltac simplify := simpl in *; clear_obvious; simpl in *; try discriminate. - Ltac rt := repeat match goal with | [ _ : error _ _ = OK _ _ _ |- _ ] => discriminate @@ -473,8 +465,8 @@ Section CORRECTNESS. | [ _ : context[match ?x with | _ => _ end] |- _ ] => destruct x end. - - - destruct c, chunk, addr, args; simplify; rt; simplify. + - (* FIXME: Should be able to use the spec to avoid destructing here. *) + destruct c, chunk, addr, args; simplify; rt; simplify. + admit. (* FIXME: Alignment *) + admit. @@ -486,104 +478,75 @@ Section CORRECTNESS. destruct (Archi.ptr64) eqn:ARCHI; simplify. rewrite ARCHI in H0. simplify. - (* Out of bounds reads are undefined behaviour. *) - assert (forall ptr v, f.(RTL.fn_stacksize) <= ptr -> - Mem.loadv AST.Mint32 m (Values.Val.offset_ptr sp (Integers.Ptrofs.repr ptr)) = Some v -> - v = Values.Vundef) as INVALID_READ by admit. - - (* Case split on whether read is out of bounds. *) - destruct (RTL.fn_stacksize f <=? Integers.Ptrofs.unsigned i0) eqn:BOUND. - * assert (RTL.fn_stacksize f <= Integers.Ptrofs.unsigned i0) as OUT_OF_BOUNDS. { - apply Zle_bool_imp_le. assumption. - } - eapply INVALID_READ in OUT_OF_BOUNDS. - 2: { rewrite Integers.Ptrofs.repr_unsigned. eassumption. } - - eexists. split. - eapply Smallstep.plus_one. - eapply HTL.step_module; eauto. - econstructor. econstructor. econstructor. simplify. - econstructor. econstructor. econstructor. econstructor. simplify. - - unfold Verilog.arr_assocmap_lookup. rewrite H3. - reflexivity. - - simplify. unfold_merge. apply AssocMap.gss. - - simplify. rewrite assumption_32bit. - constructor. - - unfold_merge. - apply regs_lessdef_add_greater. - apply greater_than_max_func. - - apply regs_lessdef_add_match. - rewrite OUT_OF_BOUNDS. - constructor. assumption. - assumption. + (** Here we are assuming that any stack read will be within the bounds + of the activation record. *) + assert (0 <= Integers.Ptrofs.unsigned i0) as READ_BOUND_LOW by admit. + assert (Integers.Ptrofs.unsigned i0 < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH by admit. - unfold state_st_wf. inversion 1. simplify. - unfold_merge. apply AssocMap.gss. - - assumption. - - econstructor; simplify; try reflexivity; eassumption. - - * assert (Integers.Ptrofs.unsigned i0 < f.(RTL.fn_stacksize)) as IN_BOUNDS. { - rewrite Z.leb_antisym in BOUND. - rewrite negb_false_iff in BOUND. - apply Zlt_is_lt_bool. assumption. - } + eexists. split. + eapply Smallstep.plus_one. + eapply HTL.step_module; eauto. + econstructor. econstructor. econstructor. simplify. + econstructor. econstructor. econstructor. econstructor. simplify. - eexists. split. - eapply Smallstep.plus_one. - eapply HTL.step_module; eauto. - econstructor. econstructor. econstructor. simplify. - econstructor. econstructor. econstructor. econstructor. simplify. + unfold Verilog.arr_assocmap_lookup. setoid_rewrite H3. + f_equal. - unfold Verilog.arr_assocmap_lookup. rewrite H3. - reflexivity. + simplify. unfold Verilog.merge_regs. + unfold_merge. + rewrite AssocMap.gso. + apply AssocMap.gss. + apply st_greater_than_res. - simplify. unfold_merge. apply AssocMap.gss. + simplify. rewrite assumption_32bit. + constructor. - simplify. rewrite assumption_32bit. - constructor. + unfold Verilog.merge_regs. unfold_merge. + apply regs_lessdef_add_match. - unfold_merge. - apply regs_lessdef_add_greater. - apply greater_than_max_func. + (** Equality proof *) + admit. - apply regs_lessdef_add_match. + apply regs_lessdef_add_greater. + apply greater_than_max_func. + assumption. assumption. - assert (exists ptr, (Z.to_nat ptr / 4)%nat = (valueToNat (ZToValue 32 (Integers.Ptrofs.unsigned i0 / 4)))) - by admit. - simplify. rewrite <- H5. - specialize (H4 x). - assert (0 <= x < Z.of_nat (Z.to_nat (RTL.fn_stacksize f / 4))) by admit. - apply H4 in H0. - invert H0. - assert (Integers.Ptrofs.repr x = i0) by admit. - rewrite H0 in H6. - rewrite H1 in H6. - invert H6. - assumption. + unfold state_st_wf. inversion 1. simplify. + unfold Verilog.merge_regs. + unfold_merge. rewrite AssocMap.gso. + apply AssocMap.gss. + apply st_greater_than_res. - assert (Integers.Ptrofs.repr x = i0) by admit. - rewrite H0 in H7. - rewrite H1 in H7. - discriminate. + assumption. - assumption. + econstructor. + repeat split; simplify. + unfold HTL.empty_stack. + simplify. + unfold Verilog.merge_arrs. + + rewrite AssocMap.gcombine. + 2: { reflexivity. } + rewrite AssocMap.gss. + unfold Verilog.merge_arr. + setoid_rewrite H3. + reflexivity. - assumption. + rewrite combine_length. + unfold arr_repeat. simplify. + rewrite list_repeat_len. + reflexivity. - unfold state_st_wf. inversion 1. simplify. - unfold_merge. apply AssocMap.gss. + unfold arr_repeat. simplify. + rewrite list_repeat_len. + congruence. - assumption. + intros. + erewrite array_get_error_equal. + eauto. apply combine_none. - econstructor; simplify; try reflexivity; eassumption. + assumption. - destruct c, chunk, addr, args; simplify; rt; simplify. + admit. -- cgit From f5172e5c66ab7175d5e90acee69e88ac214f4b0f Mon Sep 17 00:00:00 2001 From: James Pollard Date: Thu, 18 Jun 2020 01:34:49 +0100 Subject: Finish AInStack proof with minor assertions. --- src/translation/HTLgenproof.v | 53 +++++++++++++++++-------------------------- 1 file changed, 21 insertions(+), 32 deletions(-) diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index 8a246f3..f37fc8d 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -52,9 +52,9 @@ Inductive match_arrs (m : HTL.module) (f : RTL.function) (sp : Values.val) (mem (forall ptr, 0 <= ptr < Z.of_nat m.(HTL.mod_stk_len) -> opt_val_value_lessdef (Mem.loadv AST.Mint32 mem - (Values.Val.offset_ptr sp (Integers.Ptrofs.repr ptr))) + (Values.Val.offset_ptr sp (Integers.Ptrofs.repr (4 * ptr)))) (Option.default (NToValue 32 0) - (Option.join (array_get_error (Z.to_nat ptr / 4) stack)))) -> + (Option.join (array_get_error (Z.to_nat ptr) stack)))) -> match_arrs m f sp mem asa. Inductive match_stacks (mem : mem) : list RTL.stackframe -> list HTL.stackframe -> Prop := @@ -465,6 +465,8 @@ Section CORRECTNESS. | [ _ : context[match ?x with | _ => _ end] |- _ ] => destruct x end. + Opaque Z.mul. + - (* FIXME: Should be able to use the spec to avoid destructing here. *) destruct c, chunk, addr, args; simplify; rt; simplify. @@ -478,7 +480,6 @@ Section CORRECTNESS. destruct (Archi.ptr64) eqn:ARCHI; simplify. rewrite ARCHI in H0. simplify. - (** Here we are assuming that any stack read will be within the bounds of the activation record. *) assert (0 <= Integers.Ptrofs.unsigned i0) as READ_BOUND_LOW by admit. @@ -506,7 +507,18 @@ Section CORRECTNESS. apply regs_lessdef_add_match. (** Equality proof *) - admit. + assert (forall x, valueToNat (ZToValue 32 x) = Z.to_nat x) by admit. (* FIXME: 32-bit issue. *) + (* Modular arithmetic facts. *) + assert (0 <= Integers.Ptrofs.unsigned i0 / 4 < RTL.fn_stacksize f / 4) by admit. + assert (4 * (Integers.Ptrofs.unsigned i0 / 4) = Integers.Ptrofs.unsigned i0) by admit. + rewrite H0. + specialize (H5 (Integers.Ptrofs.unsigned i0 / 4)). + rewrite Z2Nat.id in H5; try lia. + apply H5 in H4. + rewrite H6 in H4. + rewrite Integers.Ptrofs.repr_unsigned in H4. + rewrite H1 in H4. + invert H4. assumption. apply regs_lessdef_add_greater. apply greater_than_max_func. @@ -553,34 +565,11 @@ Section CORRECTNESS. + admit. + admit. - + eexists. split. - eapply Smallstep.plus_one. - eapply HTL.step_module; eauto. - econstructor. econstructor. econstructor. simplify. - eapply Verilog.stmnt_runp_Vblock_arr. simplify. - econstructor. econstructor. econstructor. simplify. - - reflexivity. - - unfold_merge. apply AssocMap.gss. - - simplify. rewrite assumption_32bit. econstructor. - - unfold_merge. - apply regs_lessdef_add_greater. - apply greater_than_max_func. - assumption. - - assumption. - - unfold state_st_wf. inversion 1. simplify. - unfold_merge. apply AssocMap.gss. - - admit. - - econstructor; simplify; try reflexivity. - admit. - + (* + eexists. split. *) + (* eapply Smallstep.plus_one. *) + (* eapply HTL.step_module; eauto. *) + (* econstructor. econstructor. econstructor. simplify. *) + + admit. - eexists. split. apply Smallstep.plus_one. eapply HTL.step_module; eauto. -- cgit From 6b20fbbeaad23724ca7fbcc10c9445f5cb94b699 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Thu, 18 Jun 2020 17:12:20 +0100 Subject: Tidy up proof. --- src/translation/HTLgenproof.v | 27 +++++++++++++++------------ src/verilog/Array.v | 2 +- 2 files changed, 16 insertions(+), 13 deletions(-) diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index f37fc8d..a5396af 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -482,8 +482,8 @@ Section CORRECTNESS. (** Here we are assuming that any stack read will be within the bounds of the activation record. *) - assert (0 <= Integers.Ptrofs.unsigned i0) as READ_BOUND_LOW by admit. - assert (Integers.Ptrofs.unsigned i0 < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH by admit. + assert (0 <= Integers.Ptrofs.unsigned i0 / 4) as READ_BOUND_LOW by admit. + assert (Integers.Ptrofs.unsigned i0 / 4 < f.(RTL.fn_stacksize) / 4) as READ_BOUND_HIGH by admit. eexists. split. eapply Smallstep.plus_one. @@ -507,18 +507,21 @@ Section CORRECTNESS. apply regs_lessdef_add_match. (** Equality proof *) - assert (forall x, valueToNat (ZToValue 32 x) = Z.to_nat x) by admit. (* FIXME: 32-bit issue. *) - (* Modular arithmetic facts. *) - assert (0 <= Integers.Ptrofs.unsigned i0 / 4 < RTL.fn_stacksize f / 4) by admit. - assert (4 * (Integers.Ptrofs.unsigned i0 / 4) = Integers.Ptrofs.unsigned i0) by admit. - rewrite H0. + (* FIXME: 32-bit issue. *) + assert (forall x, valueToNat (ZToValue 32 x) = Z.to_nat x) as VALUE_IDENTITY by admit. + assert (4 * (Integers.Ptrofs.unsigned i0 / 4) = Integers.Ptrofs.unsigned i0) as MOD_IDENTITY. + { + rewrite Z.mul_comm. + rewrite ZLib.div_mul_undo; lia. + } + rewrite VALUE_IDENTITY. specialize (H5 (Integers.Ptrofs.unsigned i0 / 4)). rewrite Z2Nat.id in H5; try lia. - apply H5 in H4. - rewrite H6 in H4. - rewrite Integers.Ptrofs.repr_unsigned in H4. - rewrite H1 in H4. - invert H4. assumption. + exploit H5; auto; intros. + rewrite MOD_IDENTITY in H0. + rewrite Integers.Ptrofs.repr_unsigned in H0. + rewrite H1 in H0. + invert H0. assumption. apply regs_lessdef_add_greater. apply greater_than_max_func. diff --git a/src/verilog/Array.v b/src/verilog/Array.v index 705dea7..fc52f04 100644 --- a/src/verilog/Array.v +++ b/src/verilog/Array.v @@ -104,7 +104,7 @@ Proof. apply not_iff_compat in H1. apply <- H1 in H0. - destruct (nth_error (arr_contents a) i ) eqn:EQ; try contradiction; eauto. + destruct (nth_error (arr_contents a) i) eqn:EQ; try contradiction; eauto. Qed. Lemma array_get_error_set_bound {A : Type} : -- cgit From b59a2e2913aa7ad010c0652e909ae790c07c7281 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Thu, 18 Jun 2020 17:47:21 +0100 Subject: Enforce stack size alignment to fix proof. --- src/translation/HTLgen.v | 4 +++- src/translation/HTLgenproof.v | 25 +++++++++++++++++-------- src/translation/HTLgenspec.v | 3 ++- 3 files changed, 22 insertions(+), 10 deletions(-) diff --git a/src/translation/HTLgen.v b/src/translation/HTLgen.v index 73f2b63..f661aa6 100644 --- a/src/translation/HTLgen.v +++ b/src/translation/HTLgen.v @@ -436,6 +436,7 @@ Definition create_arr (i : option io) (sz : nat) (ln : nat) : mon (reg * nat) := (create_arr_state_incr s sz ln i). Definition transf_module (f: function) : mon module := + if (Z.eq_dec (Z.modulo f.(fn_stacksize) 4) 0) then do fin <- create_reg (Some Voutput) 1; do rtrn <- create_reg (Some Voutput) 32; do (stack, stack_len) <- create_arr None 32 (Z.to_nat (f.(fn_stacksize) / 4)); @@ -459,7 +460,8 @@ Definition transf_module (f: function) : mon module := rst clk current_state.(st_scldecls) - current_state.(st_arrdecls)). + current_state.(st_arrdecls)) + else error (Errors.msg "Stack size misalignment."). Definition max_state (f: function) : state := let st := Pos.succ (max_reg_function f) in diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index a5396af..77a11dc 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -482,8 +482,8 @@ Section CORRECTNESS. (** Here we are assuming that any stack read will be within the bounds of the activation record. *) - assert (0 <= Integers.Ptrofs.unsigned i0 / 4) as READ_BOUND_LOW by admit. - assert (Integers.Ptrofs.unsigned i0 / 4 < f.(RTL.fn_stacksize) / 4) as READ_BOUND_HIGH by admit. + assert (0 <= Integers.Ptrofs.unsigned i0) as READ_BOUND_LOW by admit. + assert (Integers.Ptrofs.unsigned i0 < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH by admit. eexists. split. eapply Smallstep.plus_one. @@ -509,16 +509,25 @@ Section CORRECTNESS. (** Equality proof *) (* FIXME: 32-bit issue. *) assert (forall x, valueToNat (ZToValue 32 x) = Z.to_nat x) as VALUE_IDENTITY by admit. - assert (4 * (Integers.Ptrofs.unsigned i0 / 4) = Integers.Ptrofs.unsigned i0) as MOD_IDENTITY. - { - rewrite Z.mul_comm. - rewrite ZLib.div_mul_undo; lia. - } rewrite VALUE_IDENTITY. specialize (H5 (Integers.Ptrofs.unsigned i0 / 4)). rewrite Z2Nat.id in H5; try lia. exploit H5; auto; intros. - rewrite MOD_IDENTITY in H0. + 1: { + split. + - apply Z.div_pos; lia. + - apply Zmult_lt_reg_r with (p := 4); try lia. + repeat rewrite ZLib.div_mul_undo; lia. + } + 2: { + assert (0 < RTL.fn_stacksize f) by lia. + apply Z.div_pos; lia. + } + replace (4 * (Integers.Ptrofs.unsigned i0 / 4)) with (Integers.Ptrofs.unsigned i0) in H0. + 2: { + rewrite Z.mul_comm. + rewrite ZLib.div_mul_undo; lia. + } rewrite Integers.Ptrofs.repr_unsigned in H0. rewrite H1 in H0. invert H0. assumption. diff --git a/src/translation/HTLgenspec.v b/src/translation/HTLgenspec.v index 1a9144c..a916cb5 100644 --- a/src/translation/HTLgenspec.v +++ b/src/translation/HTLgenspec.v @@ -188,6 +188,7 @@ Inductive tr_module (f : RTL.function) : module -> Prop := (forall pc i, Maps.PTree.get pc f.(RTL.fn_code) = Some i -> tr_code f.(RTL.fn_code) pc i data control fin rtrn st stk) -> stk_len = Z.to_nat (f.(RTL.fn_stacksize) / 4) -> + Z.modulo (f.(RTL.fn_stacksize)) 4 = 0 -> m = (mkmodule f.(RTL.fn_params) data control @@ -451,7 +452,7 @@ Proof. inversion s; subst. unfold transf_module in *. - monadInv Heqr. + destruct (Z.eq_dec (RTL.fn_stacksize f mod 4) 0); monadInv Heqr. (* TODO: We should be able to fold this into the automation. *) pose proof (create_arr_inv _ _ _ _ _ _ _ _ EQ0) as STK_LEN. -- cgit From d216c3b6dfbd80f49296b47ba46d18603c723804 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Fri, 19 Jun 2020 00:35:41 +0100 Subject: Working on proof. --- src/translation/HTLgenproof.v | 96 ++++++++++++++++++++++++++++++++++++++----- 1 file changed, 86 insertions(+), 10 deletions(-) diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index 77a11dc..046ae06 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -69,13 +69,31 @@ Inductive match_stacks (mem : mem) : list RTL.stackframe -> list HTL.stackframe match_stacks mem (RTL.Stackframe r f sp pc rs :: cs) (HTL.Stackframe r m pc asr asa :: lr). +Definition reg_stack_based_pointers (sp : Values.block) (rs : Registers.Regmap.t Values.val) : Prop := + forall r, match Registers.Regmap.get r rs with + | Values.Vptr sp' off => sp' = sp + | _ => True + end. + +Definition arr_stack_based_pointers (spb : Values.block) (m : mem) (stack_length : Z) + (sp : Values.val) : Prop := + forall ptr, + 0 <= ptr < (stack_length / 4) -> + match Mem.loadv AST.Mint32 m (Values.Val.offset_ptr sp (Integers.Ptrofs.repr (4 * ptr))) with + | Some (Values.Vptr sp' off) => sp' = spb + | _ => True + end. + Inductive match_states : RTL.state -> HTL.state -> Prop := -| match_state : forall asa asr sf f sp rs mem m st res +| match_state : forall asa asr sf f sp sp' rs mem m st res (MASSOC : match_assocmaps f rs asr) (TF : tr_module f m) (WF : state_st_wf m (HTL.State res m st asr asa)) (MS : match_stacks mem sf res) - (MARR : match_arrs m f sp mem asa), + (MARR : match_arrs m f sp mem asa) + (SP : sp = Values.Vptr sp' (Integers.Ptrofs.repr 0)) + (RSBP: reg_stack_based_pointers sp' rs) + (ASBP: arr_stack_based_pointers sp' mem (f.(RTL.fn_stacksize)) sp), match_states (RTL.State sf f sp st rs mem) (HTL.State res m st asr asa) | match_returnstate : @@ -339,7 +357,7 @@ Section CORRECTNESS. (* prove match_state *) rewrite assumption_32bit. - constructor; auto; simplify. + econstructor; simplify; eauto. unfold Verilog.merge_regs. unfold_merge. simpl. apply regs_lessdef_add_greater. apply greater_than_max_func. @@ -471,8 +489,67 @@ Section CORRECTNESS. destruct c, chunk, addr, args; simplify; rt; simplify. + admit. (* FIXME: Alignment *) - + admit. - + admit. + + + (* FIXME: Why is this degenerate? Should we support this mode? *) + unfold Op.eval_addressing in H0. + destruct (Archi.ptr64) eqn:ARCHI; simplify. + destruct (Registers.Regmap.get r0 rs) eqn:EQr0; discriminate. + + + invert MARR. simplify. + + unfold Op.eval_addressing in H0. + destruct (Archi.ptr64) eqn:ARCHI; simplify. + + unfold reg_stack_based_pointers in RSBP. + pose proof (RSBP r0) as RSBPr0. + pose proof (RSBP r1) as RSBPr1. + + destruct (Registers.Regmap.get r0 rs) eqn:EQr0; + destruct (Registers.Regmap.get r1 rs) eqn:EQr1; simplify. + + rewrite ARCHI in H1. simplify. + subst. + + eexists. split. + eapply Smallstep.plus_one. + eapply HTL.step_module; eauto. + econstructor. econstructor. econstructor. simplify. + econstructor. econstructor. econstructor. econstructor. simplify. + econstructor. + econstructor. + econstructor. + econstructor. + econstructor. + econstructor. + econstructor. + econstructor. + econstructor. + econstructor. + econstructor. + econstructor. + econstructor. + econstructor. + econstructor. + econstructor. + econstructor. + econstructor. (* FIXME: Oh dear. *) + + unfold Verilog.arr_assocmap_lookup. simplify. setoid_rewrite H3. + f_equal. + + simplify. unfold Verilog.merge_regs. + unfold_merge. + rewrite AssocMap.gso. + apply AssocMap.gss. + apply st_greater_than_res. + + simplify. rewrite assumption_32bit. + econstructor; simplify; eauto. + + unfold Verilog.merge_regs. unfold_merge. + apply regs_lessdef_add_match. + + all: admit. + invert MARR. simplify. @@ -501,7 +578,7 @@ Section CORRECTNESS. apply st_greater_than_res. simplify. rewrite assumption_32bit. - constructor. + econstructor; simplify; eauto. unfold Verilog.merge_regs. unfold_merge. apply regs_lessdef_add_match. @@ -534,7 +611,7 @@ Section CORRECTNESS. apply regs_lessdef_add_greater. apply greater_than_max_func. - assumption. assumption. + assumption. unfold state_st_wf. inversion 1. simplify. unfold Verilog.merge_regs. @@ -542,8 +619,6 @@ Section CORRECTNESS. apply AssocMap.gss. apply st_greater_than_res. - assumption. - econstructor. repeat split; simplify. unfold HTL.empty_stack. @@ -569,9 +644,10 @@ Section CORRECTNESS. intros. erewrite array_get_error_equal. eauto. apply combine_none. - assumption. + (* FIXME: RSBP Proof. *) + - destruct c, chunk, addr, args; simplify; rt; simplify. + admit. + admit. -- cgit From 11ff840afe29c5340582e513613dc70c13879997 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sat, 20 Jun 2020 10:18:25 +0100 Subject: Add proof of nat equiv --- src/verilog/Value.v | 17 +++++++++++++++-- 1 file changed, 15 insertions(+), 2 deletions(-) diff --git a/src/verilog/Value.v b/src/verilog/Value.v index b1ee353..18a69ef 100644 --- a/src/verilog/Value.v +++ b/src/verilog/Value.v @@ -19,7 +19,7 @@ (* begin hide *) From bbv Require Import Word. From bbv Require HexNotation WordScope. -From Coq Require Import ZArith.ZArith FSets.FMapPositive. +From Coq Require Import ZArith.ZArith FSets.FMapPositive Lia. From compcert Require Import lib.Integers common.Values. (* end hide *) @@ -335,7 +335,7 @@ Proof. rewrite uvalueToZ_ZToValue. auto. rewrite positive_nat_Z. split. apply Zle_0_pos. - assert (p < 2 ^ (Pos.size p))%positive. apply Pos.size_gt. + assert (p < 2 ^ (Pos.size p))%positive by apply Pos.size_gt. inversion H. rewrite <- Z.compare_lt_iff. rewrite <- H1. simpl. rewrite <- Pos2Z.inj_pow_pos. trivial. Qed. @@ -366,3 +366,16 @@ Lemma boolToValue_ValueToBool : forall b, valueToBool (boolToValue 32 b) = b. Proof. destruct b; unfold valueToBool, boolToValue; simpl; trivial. Qed. + +Lemma ZToValue_valueToNat : + forall x sz, + (x < 2^(Z.of_nat sz))%Z -> + valueToNat (ZToValue sz x) = Z.to_nat x. +Proof. + destruct x; intros; unfold ZToValue, valueToNat; simpl. + - rewrite wzero'_def. apply wordToNat_wzero. + - rewrite posToWord_nat. rewrite wordToNat_natToWord_2. auto. + inversion H. destruct (2 ^ Z.of_nat sz)%Z eqn:?; try discriminate. + Set Printing All. + Search positive Z. + - lia. -- cgit From ec9b22d01cd89aecb95da02067919423a0f1f884 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Sat, 20 Jun 2020 17:47:52 +0100 Subject: Finish structure of Aindexed2scaled ILoad proof. --- src/translation/HTLgenproof.v | 186 ++++++++++++++++++++++++++++++++++++++---- src/verilog/Value.v | 5 ++ 2 files changed, 177 insertions(+), 14 deletions(-) diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index 046ae06..be7538c 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -69,20 +69,24 @@ Inductive match_stacks (mem : mem) : list RTL.stackframe -> list HTL.stackframe match_stacks mem (RTL.Stackframe r f sp pc rs :: cs) (HTL.Stackframe r m pc asr asa :: lr). +Definition stack_based (v : Values.val) (sp : Values.block) : Prop := + match v with + | Values.Vptr sp' off => sp' = sp + | _ => True + end. + Definition reg_stack_based_pointers (sp : Values.block) (rs : Registers.Regmap.t Values.val) : Prop := - forall r, match Registers.Regmap.get r rs with - | Values.Vptr sp' off => sp' = sp - | _ => True - end. + forall r, stack_based (Registers.Regmap.get r rs) sp. Definition arr_stack_based_pointers (spb : Values.block) (m : mem) (stack_length : Z) (sp : Values.val) : Prop := forall ptr, 0 <= ptr < (stack_length / 4) -> - match Mem.loadv AST.Mint32 m (Values.Val.offset_ptr sp (Integers.Ptrofs.repr (4 * ptr))) with - | Some (Values.Vptr sp' off) => sp' = spb - | _ => True - end. + stack_based (Option.default + Values.Vundef + (Mem.loadv AST.Mint32 m + (Values.Val.offset_ptr sp (Integers.Ptrofs.repr (4 * ptr))))) + spb. Inductive match_states : RTL.state -> HTL.state -> Prop := | match_state : forall asa asr sf f sp sp' rs mem m st res @@ -514,10 +518,10 @@ Section CORRECTNESS. eapply Smallstep.plus_one. eapply HTL.step_module; eauto. econstructor. econstructor. econstructor. simplify. - econstructor. econstructor. econstructor. econstructor. simplify. - econstructor. - econstructor. - econstructor. + econstructor. econstructor. econstructor. simplify. + eapply Verilog.erun_Vbinop with (EQ := ?[EQ1]). (* FIXME: These will be shelved and cause sadness. *) + eapply Verilog.erun_Vbinop with (EQ := ?[EQ2]). + eapply Verilog.erun_Vbinop with (EQ := ?[EQ3]). econstructor. econstructor. econstructor. @@ -526,6 +530,7 @@ Section CORRECTNESS. econstructor. econstructor. econstructor. + eapply Verilog.erun_Vbinop with (EQ := ?[EQ4]). econstructor. econstructor. econstructor. @@ -549,7 +554,135 @@ Section CORRECTNESS. unfold Verilog.merge_regs. unfold_merge. apply regs_lessdef_add_match. - all: admit. + pose proof H1. + assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) by reflexivity. + rewrite H4 in H5. + setoid_rewrite Integers.Ptrofs.add_zero_l in H5. + + specialize (H5 (uvalueToZ + (vplus (vplus (vdivs asr # r0 (ZToValue 32 4) ?EQ3) (ZToValue 32 (z0 / 4)) ?EQ2) + (vmul asr # r1 (ZToValue 32 (z / 4)) ?EQ4) ?EQ1))). + + assert (0 <= uvalueToZ + (vplus (vplus (vdivs asr # r0 (ZToValue 32 4) ?EQ3) (ZToValue 32 (z0 / 4)) ?EQ2) + (vmul asr # r1 (ZToValue 32 (z / 4)) ?EQ4) ?EQ1) < + Z.of_nat (Z.to_nat (RTL.fn_stacksize f / 4))) by admit. + apply H5 in H6. + + invert MASSOC. + pose proof (H7 r0). + pose proof (H7 r1). + assert (HPler0 : Ple r0 (RTL.max_reg_function f)) + by (eapply RTL.max_reg_function_use; eauto; simplify; eauto). + assert (HPler1 : Ple r1 (RTL.max_reg_function f)) + by (eapply RTL.max_reg_function_use; eauto; simpl; auto). + apply H8 in HPler0. + apply H10 in HPler1. + invert HPler0; invert HPler1; try congruence. + rewrite EQr0 in H12. + rewrite EQr1 in H13. + + assert ((Integers.Ptrofs.unsigned + (Integers.Ptrofs.add i0 + (Integers.Ptrofs.of_int + (Integers.Int.add (Integers.Int.mul i1 (Integers.Int.repr z)) + (Integers.Int.repr z0))))) = + (Integers.Ptrofs.unsigned + (Integers.Ptrofs.repr + (4 * + uvalueToZ + (vplus + (vplus (vdivs asr # r0 (ZToValue 32 4) ?EQ3) + (ZToValue 32 (z0 / 4)) ?EQ2) + (vmul asr # r1 (ZToValue 32 (z / 4)) ?EQ4) + ?EQ1))))) by admit. + + rewrite <- H19 in H6. + rewrite H0 in H6. + invert H6. + assert (forall x, Z.to_nat (uvalueToZ x) = valueToNat x) as VALUE_IDENTITY by admit. + rewrite VALUE_IDENTITY in H21. + assumption. + + apply regs_lessdef_add_greater. + apply greater_than_max_func. + assumption. + + unfold state_st_wf. inversion 1. simplify. + unfold Verilog.merge_regs. + unfold_merge. rewrite AssocMap.gso. + apply AssocMap.gss. + apply st_greater_than_res. + + econstructor. + repeat split; simplify. + unfold HTL.empty_stack. + simplify. + unfold Verilog.merge_arrs. + + rewrite AssocMap.gcombine. + 2: { reflexivity. } + rewrite AssocMap.gss. + unfold Verilog.merge_arr. + setoid_rewrite H3. + reflexivity. + + rewrite combine_length. + unfold arr_repeat. simplify. + rewrite list_repeat_len. + reflexivity. + + unfold arr_repeat. simplify. + rewrite list_repeat_len. + congruence. + + intros. + erewrite array_get_error_equal. + eauto. apply combine_none. + assumption. + + unfold reg_stack_based_pointers. intros. + destruct (Pos.eq_dec r2 dst); subst. + + rewrite Registers.Regmap.gss. + + unfold arr_stack_based_pointers in ASBP. + specialize (ASBP ((Integers.Ptrofs.unsigned + (Integers.Ptrofs.add i0 + (Integers.Ptrofs.of_int + (Integers.Int.add (Integers.Int.mul i1 (Integers.Int.repr z)) + (Integers.Int.repr z0))))) / 4)). + exploit ASBP; auto; intros. + 1: { + split. + - admit. (*apply Z.div_pos; lia.*) + - admit. (* apply Zmult_lt_reg_r with (p := 4); try lia. *) + (* repeat rewrite ZLib.div_mul_undo; lia. *) + } + replace (4 * ((Integers.Ptrofs.unsigned + (Integers.Ptrofs.add i0 + (Integers.Ptrofs.of_int + (Integers.Int.add (Integers.Int.mul i1 (Integers.Int.repr z)) + (Integers.Int.repr z0))))) / 4)) with (Integers.Ptrofs.unsigned + (Integers.Ptrofs.add i0 + (Integers.Ptrofs.of_int + (Integers.Int.add (Integers.Int.mul i1 (Integers.Int.repr z)) + (Integers.Int.repr z0))))) in H0. + 2: { + rewrite Z.mul_comm. + admit. + (* rewrite ZLib.div_mul_undo; lia. *) + } + rewrite Integers.Ptrofs.repr_unsigned in H0. + simplify. + assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) by reflexivity. + rewrite H4 in H0. + setoid_rewrite Integers.Ptrofs.add_zero_l in H0. + rewrite H1 in H0. + simplify. + assumption. + + rewrite Registers.Regmap.gso; auto. + invert MARR. simplify. @@ -646,7 +779,32 @@ Section CORRECTNESS. eauto. apply combine_none. assumption. - (* FIXME: RSBP Proof. *) + unfold reg_stack_based_pointers. intros. + destruct (Pos.eq_dec r0 dst); subst. + + rewrite Registers.Regmap.gss. + + unfold arr_stack_based_pointers in ASBP. + specialize (ASBP (Integers.Ptrofs.unsigned i0 / 4)). + exploit ASBP; auto; intros. + 1: { + split. + - apply Z.div_pos; lia. + - apply Zmult_lt_reg_r with (p := 4); try lia. + repeat rewrite ZLib.div_mul_undo; lia. + } + replace (4 * (Integers.Ptrofs.unsigned i0 / 4)) with (Integers.Ptrofs.unsigned i0) in H0. + 2: { + rewrite Z.mul_comm. + rewrite ZLib.div_mul_undo; lia. + } + rewrite Integers.Ptrofs.repr_unsigned in H0. + simplify. + rewrite H1 in H0. + simplify. + assumption. + + rewrite Registers.Regmap.gso; auto. - destruct c, chunk, addr, args; simplify; rt; simplify. + admit. diff --git a/src/verilog/Value.v b/src/verilog/Value.v index b1ee353..818f625 100644 --- a/src/verilog/Value.v +++ b/src/verilog/Value.v @@ -294,6 +294,11 @@ Inductive val_value_lessdef: val -> value -> Prop := forall i v', i = valueToInt v' -> val_value_lessdef (Vint i) v' +| val_value_lessdef_ptr: + forall b off v', + off = Ptrofs.repr (valueToZ v') -> + (Z.modulo (valueToZ v') 4) = 0%Z -> + val_value_lessdef (Vptr b off) v' | lessdef_undef: forall v, val_value_lessdef Vundef v. Inductive opt_val_value_lessdef: option val -> value -> Prop := -- cgit From 2b4a5fcb4b58122d9bf8ae52f03c9855ffeb1d77 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Sun, 21 Jun 2020 13:47:32 +0100 Subject: Lea op now checks alignment. --- src/translation/HTLgen.v | 27 +++++++++++++++------------ 1 file changed, 15 insertions(+), 12 deletions(-) diff --git a/src/translation/HTLgen.v b/src/translation/HTLgen.v index f661aa6..1c4130b 100644 --- a/src/translation/HTLgen.v +++ b/src/translation/HTLgen.v @@ -246,18 +246,21 @@ Definition translate_condition (c : Op.condition) (args : list reg) : mon expr : end. Definition translate_eff_addressing (a: Op.addressing) (args: list reg) : mon expr := - match a, args with + match a, args with (* TODO: We should be more methodical here; what are the possibilities?*) | Op.Aindexed off, r1::nil => ret (boplitz Vadd r1 off) - | Op.Aindexed2 off, r1::r2::nil => ret (Vbinop Vadd (Vvar r1) (boplitz Vadd r2 off)) | Op.Ascaled scale offset, r1::nil => - ret (Vbinop Vadd (boplitz Vmul r1 scale) (Vlit (ZToValue 32 offset))) - | Op.Aindexed2scaled scale offset, r1::r2::nil => - ret (Vbinop Vadd (boplitz Vadd r1 offset) (boplitz Vmul r2 scale)) - (* Stack arrays/referenced variables *) - | Op.Ainstack a, nil => (* We need to be sure that the base address is aligned *) - let a := Integers.Ptrofs.unsigned a in (* FIXME: Assuming stack offsets are +ve; is this ok? *) - ret (Vlit (ZToValue 32 a)) - | _, _ => error (Errors.msg "Veriloggen: eff_addressing instruction not implemented: other") + if ((Z.eqb (Z.modulo scale 4) 0) && (Z.eqb (Z.modulo offset 4) 0)) + then ret (Vbinop Vadd (boplitz Vmul r1 scale) (Vlit (ZToValue 32 offset))) + else error (Errors.msg "Veriloggen: translate_eff_addressing address misaligned") + | Mint32, Op.Aindexed2scaled scale offset, r1::r2::nil => (* Typical for dynamic array addressing *) + if ((Z.eqb (Z.modulo scale 4) 0) && (Z.eqb (Z.modulo offset 4) 0)) + then ret (Vbinop Vadd (boplitz Vadd r1 offset) (boplitz Vmul r2 scale)) + else error (Errors.msg "Veriloggen: translate_eff_addressing address misaligned") + | Mint32, Op.Ainstack a, nil => (* We need to be sure that the base address is aligned *) + let a := Integers.Ptrofs.unsigned a in + if (Z.eq_dec (Z.modulo a 4) 0) then ret (Vlit (ZToValue 32 a)) + else error (Errors.msg "Veriloggen: translate_eff_addressing address misaligned") + | _, _, _ => error (Errors.msg "Veriloggen: translate_eff_addressing unsuported addressing") end. (** Translate an instruction to a statement. *) @@ -335,7 +338,7 @@ Definition add_branch_instr (e: expr) (n n1 n2: node) : mon unit := Definition translate_arr_access (mem : AST.memory_chunk) (addr : Op.addressing) (args : list reg) (stack : reg) : mon expr := match mem, addr, args with (* TODO: We should be more methodical here; what are the possibilities?*) - | Mint32, Op.Aindexed off, r1::nil => (* FIXME: Cannot guarantee alignment *) + | Mint32, Op.Aindexed off, r1::nil => ret (Vvari stack (Vbinop Vadd (boplitz Vdiv r1 4) (Vlit (ZToValue 32 (off / 4))))) | Mint32, Op.Ascaled scale offset, r1::nil => if ((Z.eqb (Z.modulo scale 4) 0) && (Z.eqb (Z.modulo offset 4) 0)) @@ -348,7 +351,7 @@ Definition translate_arr_access (mem : AST.memory_chunk) (addr : Op.addressing) (boplitz Vmul r2 (scale / 4)))) else error (Errors.msg "Veriloggen: translate_arr_access address misaligned") | Mint32, Op.Ainstack a, nil => (* We need to be sure that the base address is aligned *) - let a := Integers.Ptrofs.unsigned a in (* FIXME: Assuming stack offsets are +ve; is this ok? *) + let a := Integers.Ptrofs.unsigned a in if (Z.eq_dec (Z.modulo a 4) 0) then ret (Vvari stack (Vlit (ZToValue 32 (a / 4)))) else error (Errors.msg "Veriloggen: eff_addressing misaligned stack offset") | _, _, _ => error (Errors.msg "Veriloggen: translate_arr_access unsuported addressing") -- cgit From e05b93c540d2e0e2cb9f4ab01460eba080b65401 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Sun, 21 Jun 2020 15:01:15 +0100 Subject: Factor out addressing checks, check signed range. --- src/translation/HTLgen.v | 27 +++++++++++++++++---------- 1 file changed, 17 insertions(+), 10 deletions(-) diff --git a/src/translation/HTLgen.v b/src/translation/HTLgen.v index 1c4130b..1c2d786 100644 --- a/src/translation/HTLgen.v +++ b/src/translation/HTLgen.v @@ -17,7 +17,7 @@ *) From compcert Require Import Maps. -From compcert Require Errors Globalenvs. +From compcert Require Errors Globalenvs Integers. From compcert Require Import AST RTL. From coqup Require Import Verilog HTL Coquplib AssocMap Value Statemonad. @@ -245,22 +245,28 @@ Definition translate_condition (c : Op.condition) (args : list reg) : mon expr : | _, _ => error (Errors.msg "Veriloggen: condition instruction not implemented: other") end. +Definition check_address_parameter (p : Z) : bool := + Z.eqb (Z.modulo p 4) 0 + && Z.leb Integers.Ptrofs.min_signed p + && Z.leb p Integers.Ptrofs.min_signed. + Definition translate_eff_addressing (a: Op.addressing) (args: list reg) : mon expr := match a, args with (* TODO: We should be more methodical here; what are the possibilities?*) | Op.Aindexed off, r1::nil => ret (boplitz Vadd r1 off) | Op.Ascaled scale offset, r1::nil => - if ((Z.eqb (Z.modulo scale 4) 0) && (Z.eqb (Z.modulo offset 4) 0)) + if (check_address_parameter scale) && (check_address_parameter offset) then ret (Vbinop Vadd (boplitz Vmul r1 scale) (Vlit (ZToValue 32 offset))) else error (Errors.msg "Veriloggen: translate_eff_addressing address misaligned") - | Mint32, Op.Aindexed2scaled scale offset, r1::r2::nil => (* Typical for dynamic array addressing *) - if ((Z.eqb (Z.modulo scale 4) 0) && (Z.eqb (Z.modulo offset 4) 0)) + | Op.Aindexed2scaled scale offset, r1::r2::nil => (* Typical for dynamic array addressing *) + if (check_address_parameter scale) && (check_address_parameter offset) then ret (Vbinop Vadd (boplitz Vadd r1 offset) (boplitz Vmul r2 scale)) else error (Errors.msg "Veriloggen: translate_eff_addressing address misaligned") - | Mint32, Op.Ainstack a, nil => (* We need to be sure that the base address is aligned *) + | Op.Ainstack a, nil => (* We need to be sure that the base address is aligned *) let a := Integers.Ptrofs.unsigned a in - if (Z.eq_dec (Z.modulo a 4) 0) then ret (Vlit (ZToValue 32 a)) + if (check_address_parameter a) + then ret (Vlit (ZToValue 32 a)) else error (Errors.msg "Veriloggen: translate_eff_addressing address misaligned") - | _, _, _ => error (Errors.msg "Veriloggen: translate_eff_addressing unsuported addressing") + | _, _ => error (Errors.msg "Veriloggen: translate_eff_addressing unsuported addressing") end. (** Translate an instruction to a statement. *) @@ -341,18 +347,19 @@ Definition translate_arr_access (mem : AST.memory_chunk) (addr : Op.addressing) | Mint32, Op.Aindexed off, r1::nil => ret (Vvari stack (Vbinop Vadd (boplitz Vdiv r1 4) (Vlit (ZToValue 32 (off / 4))))) | Mint32, Op.Ascaled scale offset, r1::nil => - if ((Z.eqb (Z.modulo scale 4) 0) && (Z.eqb (Z.modulo offset 4) 0)) + if (check_address_parameter scale) && (check_address_parameter offset) then ret (Vvari stack (Vbinop Vadd (boplitz Vmul r1 (scale / 4)) (Vlit (ZToValue 32 (offset / 4))))) else error (Errors.msg "Veriloggen: translate_arr_access address misaligned") | Mint32, Op.Aindexed2scaled scale offset, r1::r2::nil => (* Typical for dynamic array addressing *) - if ((Z.eqb (Z.modulo scale 4) 0) && (Z.eqb (Z.modulo offset 4) 0)) + if (check_address_parameter scale) && (check_address_parameter offset) then ret (Vvari stack (Vbinop Vadd (Vbinop Vadd (boplitz Vdiv r1 4) (Vlit (ZToValue 32 (offset / 4)))) (boplitz Vmul r2 (scale / 4)))) else error (Errors.msg "Veriloggen: translate_arr_access address misaligned") | Mint32, Op.Ainstack a, nil => (* We need to be sure that the base address is aligned *) let a := Integers.Ptrofs.unsigned a in - if (Z.eq_dec (Z.modulo a 4) 0) then ret (Vvari stack (Vlit (ZToValue 32 (a / 4)))) + if (check_address_parameter a) + then ret (Vvari stack (Vlit (ZToValue 32 (a / 4)))) else error (Errors.msg "Veriloggen: eff_addressing misaligned stack offset") | _, _, _ => error (Errors.msg "Veriloggen: translate_arr_access unsuported addressing") end. -- cgit From be6ad3cd886b3ea79abe6addc2f2add779f55292 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Mon, 22 Jun 2020 17:15:29 +0100 Subject: Tidy up proof for Aindexed2scaled. --- src/common/Coquplib.v | 55 ++++++++- src/translation/HTLgen.v | 7 +- src/translation/HTLgenproof.v | 252 ++++++++++++++++++++++++++++-------------- src/translation/HTLgenspec.v | 7 +- 4 files changed, 231 insertions(+), 90 deletions(-) diff --git a/src/common/Coquplib.v b/src/common/Coquplib.v index efa1323..59b23ae 100644 --- a/src/common/Coquplib.v +++ b/src/common/Coquplib.v @@ -28,6 +28,7 @@ From coqup Require Import Show. (* Depend on CompCert for the basic library, as they declare and prove some useful theorems. *) From compcert.lib Require Export Coqlib. +From compcert Require Integers. Ltac unfold_rec c := unfold c; fold c. @@ -50,9 +51,61 @@ Ltac clear_obvious := | [ H : _ /\ _ |- _ ] => invert H end. -Ltac simplify := simpl in *; clear_obvious; simpl in *; try discriminate. +Ltac kill_bools := + repeat match goal with + | [ H : _ && _ = true |- _ ] => apply andb_prop in H + | [ H : _ || _ = false |- _ ] => apply orb_false_elim in H + + | [ H : _ <=? _ = true |- _ ] => apply Z.leb_le in H + | [ H : _ <=? _ = false |- _ ] => apply Z.leb_gt in H + | [ H : _ apply Z.ltb_lt in H + | [ H : _ apply Z.ltb_ge in H + | [ H : _ >=? _ = _ |- _ ] => rewrite Z.geb_leb in H + | [ H : _ >? _ = _ |- _ ] => rewrite Z.gtb_ltb in H + + | [ H : _ =? _ = true |- _ ] => apply Z.eqb_eq in H + | [ H : _ =? _ = false |- _ ] => apply Z.eqb_neq in H + end. + +Ltac unfold_constants := + repeat match goal with + | [ _ : _ |- context[Integers.Ptrofs.modulus] ] => + replace Integers.Ptrofs.modulus with 4294967296 by reflexivity + | [ H : context[Integers.Ptrofs.modulus] |- _ ] => + replace Integers.Ptrofs.modulus with 4294967296 in H by reflexivity + + | [ _ : _ |- context[Integers.Ptrofs.min_signed] ] => + replace Integers.Ptrofs.min_signed with (-2147483648) by reflexivity + | [ H : context[Integers.Ptrofs.min_signed] |- _ ] => + replace Integers.Ptrofs.min_signed with (-2147483648) in H by reflexivity + + | [ _ : _ |- context[Integers.Ptrofs.max_signed] ] => + replace Integers.Ptrofs.max_signed with 2147483647 by reflexivity + | [ H : context[Integers.Ptrofs.max_signed] |- _ ] => + replace Integers.Ptrofs.max_signed with 2147483647 in H by reflexivity + + | [ _ : _ |- context[Integers.Int.modulus] ] => + replace Integers.Int.modulus with 4294967296 by reflexivity + | [ H : context[Integers.Int.modulus] |- _ ] => + replace Integers.Int.modulus with 4294967296 in H by reflexivity + + | [ _ : _ |- context[Integers.Int.min_signed] ] => + replace Integers.Int.min_signed with (-2147483648) by reflexivity + | [ H : context[Integers.Int.min_signed] |- _ ] => + replace Integers.Int.min_signed with (-2147483648) in H by reflexivity + + | [ _ : _ |- context[Integers.Int.max_signed] ] => + replace Integers.Int.max_signed with 2147483647 by reflexivity + | [ H : context[Integers.Int.max_signed] |- _ ] => + replace Integers.Int.max_signed with 2147483647 in H by reflexivity + end. + +Ltac simplify := unfold_constants; simpl in *; + repeat (clear_obvious; kill_bools); + simpl in *; try discriminate. Global Opaque Nat.div. +Global Opaque Z.mul. (* Definition const (A B : Type) (a : A) (b : B) : A := a. diff --git a/src/translation/HTLgen.v b/src/translation/HTLgen.v index 1c2d786..2364a0f 100644 --- a/src/translation/HTLgen.v +++ b/src/translation/HTLgen.v @@ -248,7 +248,7 @@ Definition translate_condition (c : Op.condition) (args : list reg) : mon expr : Definition check_address_parameter (p : Z) : bool := Z.eqb (Z.modulo p 4) 0 && Z.leb Integers.Ptrofs.min_signed p - && Z.leb p Integers.Ptrofs.min_signed. + && Z.leb p Integers.Ptrofs.max_signed. Definition translate_eff_addressing (a: Op.addressing) (args: list reg) : mon expr := match a, args with (* TODO: We should be more methodical here; what are the possibilities?*) @@ -445,8 +445,11 @@ Definition create_arr (i : option io) (sz : nat) (ln : nat) : mon (reg * nat) := (st_controllogic s)) (create_arr_state_incr s sz ln i). +Definition stack_correct (sz : Z) : bool := + (0 <=? sz) && (Z.modulo sz 4 =? 0). + Definition transf_module (f: function) : mon module := - if (Z.eq_dec (Z.modulo f.(fn_stacksize) 4) 0) then + if stack_correct f.(fn_stacksize) then do fin <- create_reg (Some Voutput) 1; do rtrn <- create_reg (Some Voutput) 32; do (stack, stack_len) <- create_arr None 32 (Z.to_nat (f.(fn_stacksize) / 4)); diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index be7538c..21dbc73 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -18,7 +18,7 @@ From compcert Require RTL Registers AST Integers. From compcert Require Import Globalenvs Memory. -From coqup Require Import Coquplib HTLgenspec HTLgen Value AssocMap Array. +From coqup Require Import Coquplib HTLgenspec HTLgen Value AssocMap Array IntegerExtra. From coqup Require HTL Verilog. Require Import Lia. @@ -334,6 +334,7 @@ Section CORRECTNESS. exists asr' asa', HTL.step tge (HTL.State res m st asr asa) Events.E0 (HTL.State res m st asr' asa'). + Theorem transl_step_correct: forall (S1 : RTL.state) t S2, RTL.step ge S1 t S2 -> @@ -487,8 +488,6 @@ Section CORRECTNESS. | [ _ : context[match ?x with | _ => _ end] |- _ ] => destruct x end. - Opaque Z.mul. - - (* FIXME: Should be able to use the spec to avoid destructing here. *) destruct c, chunk, addr, args; simplify; rt; simplify. @@ -513,7 +512,97 @@ Section CORRECTNESS. rewrite ARCHI in H1. simplify. subst. + clear RSBPr1. + + pose proof MASSOC as MASSOC'. + invert MASSOC'. + pose proof (H0 r0). + pose proof (H0 r1). + assert (HPler0 : Ple r0 (RTL.max_reg_function f)) + by (eapply RTL.max_reg_function_use; eauto; simplify; eauto). + assert (HPler1 : Ple r1 (RTL.max_reg_function f)) + by (eapply RTL.max_reg_function_use; eauto; simpl; auto). + apply H4 in HPler0. + apply H6 in HPler1. + invert HPler0; invert HPler1; try congruence. + rewrite EQr0 in H7. + rewrite EQr1 in H8. + invert H7. invert H8. + clear H0. clear H4. clear H6. + + unfold check_address_parameter in *. simplify. + + (** Normalisation proof *) + assert + (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0)) + (Integers.Ptrofs.of_int + (Integers.Int.add (Integers.Int.mul (valueToInt asr # r1) (Integers.Int.repr z)) + (Integers.Int.repr z0))) + = + Integers.Ptrofs.repr + (4 * + Integers.Ptrofs.unsigned + (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0 / 4)) + (Integers.Ptrofs.of_int + (Integers.Int.add + (Integers.Int.mul (valueToInt asr # r1) (Integers.Int.repr (z / 4))) + (Integers.Int.repr (z0 / 4))))))) as NORMALISE. + etransitivity. + 2: { + replace (4 * + Integers.Ptrofs.unsigned + (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0 / 4)) + (Integers.Ptrofs.of_int + (Integers.Int.add + (Integers.Int.mul (valueToInt asr # r1) (Integers.Int.repr (z / 4))) + (Integers.Int.repr (z0 / 4)))))) with + (Integers.Ptrofs.unsigned (Integers.Ptrofs.repr 4) * + Integers.Ptrofs.unsigned + (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0 / 4)) + (Integers.Ptrofs.of_int + (Integers.Int.add + (Integers.Int.mul (valueToInt asr # r1) (Integers.Int.repr (z / 4))) + (Integers.Int.repr (z0 / 4)))))). + 2: { rewrite Integers.Ptrofs.unsigned_repr_eq. + rewrite Z.mod_small. reflexivity. + simplify. lia. } + + rewrite <- PtrofsExtra.mul_unsigned. + rewrite Integers.Ptrofs.mul_add_distr_r. + rewrite PtrofsExtra.mul_repr; simplify; try lia. + rewrite ZLib.div_mul_undo; try lia. + rewrite mul_of_int; simplify; try lia. + rewrite Integers.Int.mul_add_distr_r. + rewrite Integers.Int.mul_commut. + rewrite Integers.Int.mul_assoc. + rewrite IntExtra.mul_repr; simplify; try lia. + rewrite IntExtra.mul_repr; simplify; try lia. + rewrite ZLib.div_mul_undo; try lia. + rewrite Z.mul_comm. + rewrite ZLib.div_mul_undo; try lia. + reflexivity. + + split. + apply Z.div_le_lower_bound; lia. + apply Z.div_le_upper_bound; lia. + split. + apply Z.div_le_lower_bound; lia. + apply Z.div_le_upper_bound; lia. + + admit. (* FIXME: Register size 32 bits *) } + reflexivity. + + remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0)) + (Integers.Ptrofs.of_int + (Integers.Int.add (Integers.Int.mul (valueToInt asr # r1) (Integers.Int.repr z)) + (Integers.Int.repr z0)))) as OFFSET. + + (** Here we are assuming that any stack read will be within the bounds + of the activation record. *) + assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH by admit. + + (* Proof begins in earnest here. *) eexists. split. eapply Smallstep.plus_one. eapply HTL.step_module; eauto. @@ -554,55 +643,43 @@ Section CORRECTNESS. unfold Verilog.merge_regs. unfold_merge. apply regs_lessdef_add_match. - pose proof H1. + (** Equality proof *) assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) by reflexivity. - rewrite H4 in H5. + rewrite H0 in H5. clear H0. setoid_rewrite Integers.Ptrofs.add_zero_l in H5. - specialize (H5 (uvalueToZ - (vplus (vplus (vdivs asr # r0 (ZToValue 32 4) ?EQ3) (ZToValue 32 (z0 / 4)) ?EQ2) - (vmul asr # r1 (ZToValue 32 (z / 4)) ?EQ4) ?EQ1))). - - assert (0 <= uvalueToZ - (vplus (vplus (vdivs asr # r0 (ZToValue 32 4) ?EQ3) (ZToValue 32 (z0 / 4)) ?EQ2) - (vmul asr # r1 (ZToValue 32 (z / 4)) ?EQ4) ?EQ1) < - Z.of_nat (Z.to_nat (RTL.fn_stacksize f / 4))) by admit. - apply H5 in H6. + specialize (H5 + (Integers.Ptrofs.unsigned + (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0 / 4)) + (Integers.Ptrofs.of_int + (Integers.Int.add (Integers.Int.mul + (valueToInt asr # r1) + (Integers.Int.repr (z / 4))) + (Integers.Int.repr (z0 / 4))))))). - invert MASSOC. - pose proof (H7 r0). - pose proof (H7 r1). - assert (HPler0 : Ple r0 (RTL.max_reg_function f)) - by (eapply RTL.max_reg_function_use; eauto; simplify; eauto). - assert (HPler1 : Ple r1 (RTL.max_reg_function f)) - by (eapply RTL.max_reg_function_use; eauto; simpl; auto). - apply H8 in HPler0. - apply H10 in HPler1. - invert HPler0; invert HPler1; try congruence. - rewrite EQr0 in H12. - rewrite EQr1 in H13. - - assert ((Integers.Ptrofs.unsigned - (Integers.Ptrofs.add i0 - (Integers.Ptrofs.of_int - (Integers.Int.add (Integers.Int.mul i1 (Integers.Int.repr z)) - (Integers.Int.repr z0))))) = - (Integers.Ptrofs.unsigned - (Integers.Ptrofs.repr - (4 * - uvalueToZ - (vplus - (vplus (vdivs asr # r0 (ZToValue 32 4) ?EQ3) - (ZToValue 32 (z0 / 4)) ?EQ2) - (vmul asr # r1 (ZToValue 32 (z / 4)) ?EQ4) - ?EQ1))))) by admit. - - rewrite <- H19 in H6. - rewrite H0 in H6. - invert H6. - assert (forall x, Z.to_nat (uvalueToZ x) = valueToNat x) as VALUE_IDENTITY by admit. - rewrite VALUE_IDENTITY in H21. - assumption. + exploit H5; auto; intros. + rewrite Z2Nat.id. + split. + apply Integers.Ptrofs.unsigned_range_2. + admit. + apply Z_div_pos; lia. + clear H5. + + assert (Z.to_nat + (Integers.Ptrofs.unsigned + (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0 / 4)) + (Integers.Ptrofs.of_int + (Integers.Int.add + (Integers.Int.mul (valueToInt asr # r1) (Integers.Int.repr (z / 4))) + (Integers.Int.repr (z0 / 4)))))) + = + valueToNat + (vplus (vplus (vdivs asr # r0 (ZToValue 32 4) ?EQ3) (ZToValue 32 (z0 / 4)) ?EQ2) + (vmul asr # r1 (ZToValue 32 (z / 4)) ?EQ4) ?EQ1)) as EXPR_OK by admit. + rewrite <- EXPR_OK. + rewrite <- NORMALISE in H0. + rewrite H1 in H0. + invert H0. assumption. apply regs_lessdef_add_greater. apply greater_than_max_func. @@ -647,40 +724,41 @@ Section CORRECTNESS. rewrite Registers.Regmap.gss. unfold arr_stack_based_pointers in ASBP. - specialize (ASBP ((Integers.Ptrofs.unsigned - (Integers.Ptrofs.add i0 - (Integers.Ptrofs.of_int - (Integers.Int.add (Integers.Int.mul i1 (Integers.Int.repr z)) - (Integers.Int.repr z0))))) / 4)). - exploit ASBP; auto; intros. - 1: { - split. - - admit. (*apply Z.div_pos; lia.*) - - admit. (* apply Zmult_lt_reg_r with (p := 4); try lia. *) - (* repeat rewrite ZLib.div_mul_undo; lia. *) - } - replace (4 * ((Integers.Ptrofs.unsigned - (Integers.Ptrofs.add i0 - (Integers.Ptrofs.of_int - (Integers.Int.add (Integers.Int.mul i1 (Integers.Int.repr z)) - (Integers.Int.repr z0))))) / 4)) with (Integers.Ptrofs.unsigned - (Integers.Ptrofs.add i0 - (Integers.Ptrofs.of_int - (Integers.Int.add (Integers.Int.mul i1 (Integers.Int.repr z)) - (Integers.Int.repr z0))))) in H0. - 2: { - rewrite Z.mul_comm. - admit. - (* rewrite ZLib.div_mul_undo; lia. *) - } - rewrite Integers.Ptrofs.repr_unsigned in H0. - simplify. - assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) by reflexivity. - rewrite H4 in H0. - setoid_rewrite Integers.Ptrofs.add_zero_l in H0. - rewrite H1 in H0. - simplify. - assumption. + (* specialize (ASBP ((Integers.Ptrofs.unsigned *) + (* (Integers.Ptrofs.add i0 *) + (* (Integers.Ptrofs.of_int *) + (* (Integers.Int.add (Integers.Int.mul i1 (Integers.Int.repr z)) *) + (* (Integers.Int.repr z0))))) / 4)). *) + (* exploit ASBP; auto; intros. *) + (* 1: { *) + (* split. *) + (* - admit. (*apply Z.div_pos; lia.*) *) + (* - admit. (* apply Zmult_lt_reg_r with (p := 4); try lia. *) *) + (* (* repeat rewrite ZLib.div_mul_undo; lia. *) *) + (* } *) + (* replace (4 * ((Integers.Ptrofs.unsigned *) + (* (Integers.Ptrofs.add i0 *) + (* (Integers.Ptrofs.of_int *) + (* (Integers.Int.add (Integers.Int.mul i1 (Integers.Int.repr z)) *) + (* (Integers.Int.repr z0))))) / 4)) with (Integers.Ptrofs.unsigned *) + (* (Integers.Ptrofs.add i0 *) + (* (Integers.Ptrofs.of_int *) + (* (Integers.Int.add (Integers.Int.mul i1 (Integers.Int.repr z)) *) + (* (Integers.Int.repr z0))))) in H0. *) + (* 2: { *) + (* rewrite Z.mul_comm. *) + (* admit. *) + (* (* rewrite ZLib.div_mul_undo; lia. *) *) + (* } *) + (* rewrite Integers.Ptrofs.repr_unsigned in H0. *) + (* simplify. *) + (* assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) by reflexivity. *) + (* rewrite H4 in H0. *) + (* setoid_rewrite Integers.Ptrofs.add_zero_l in H0. *) + (* rewrite H1 in H0. *) + (* simplify. *) + (* assumption. *) + admit. rewrite Registers.Regmap.gso; auto. @@ -690,9 +768,10 @@ Section CORRECTNESS. destruct (Archi.ptr64) eqn:ARCHI; simplify. rewrite ARCHI in H0. simplify. + unfold check_address_parameter in EQ; simplify. + (** Here we are assuming that any stack read will be within the bounds of the activation record. *) - assert (0 <= Integers.Ptrofs.unsigned i0) as READ_BOUND_LOW by admit. assert (Integers.Ptrofs.unsigned i0 < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH by admit. eexists. split. @@ -725,12 +804,12 @@ Section CORRECTNESS. exploit H5; auto; intros. 1: { split. - - apply Z.div_pos; lia. + - apply Z.div_pos; try lia. + apply Integers.Ptrofs.unsigned_range_2. - apply Zmult_lt_reg_r with (p := 4); try lia. repeat rewrite ZLib.div_mul_undo; lia. } 2: { - assert (0 < RTL.fn_stacksize f) by lia. apply Z.div_pos; lia. } replace (4 * (Integers.Ptrofs.unsigned i0 / 4)) with (Integers.Ptrofs.unsigned i0) in H0. @@ -789,7 +868,8 @@ Section CORRECTNESS. exploit ASBP; auto; intros. 1: { split. - - apply Z.div_pos; lia. + - apply Z.div_pos; try lia. + apply Integers.Ptrofs.unsigned_range_2. - apply Zmult_lt_reg_r with (p := 4); try lia. repeat rewrite ZLib.div_mul_undo; lia. } diff --git a/src/translation/HTLgenspec.v b/src/translation/HTLgenspec.v index a916cb5..528c662 100644 --- a/src/translation/HTLgenspec.v +++ b/src/translation/HTLgenspec.v @@ -189,6 +189,7 @@ Inductive tr_module (f : RTL.function) : module -> Prop := tr_code f.(RTL.fn_code) pc i data control fin rtrn st stk) -> stk_len = Z.to_nat (f.(RTL.fn_stacksize) / 4) -> Z.modulo (f.(RTL.fn_stacksize)) 4 = 0 -> + 0 <= f.(RTL.fn_stacksize) -> m = (mkmodule f.(RTL.fn_params) data control @@ -452,7 +453,11 @@ Proof. inversion s; subst. unfold transf_module in *. - destruct (Z.eq_dec (RTL.fn_stacksize f mod 4) 0); monadInv Heqr. + unfold stack_correct in *. + destruct (0 <=? RTL.fn_stacksize f) eqn:STACK_BOUND; + destruct (RTL.fn_stacksize f mod 4 =? 0) eqn:STACK_ALIGN; + simplify; + monadInv Heqr. (* TODO: We should be able to fold this into the automation. *) pose proof (create_arr_inv _ _ _ _ _ _ _ _ EQ0) as STK_LEN. -- cgit From 3ec28d050aebc305c6df5b4b95bcf91498ff11cc Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Mon, 22 Jun 2020 17:53:54 +0100 Subject: Admit the value proof --- src/verilog/Value.v | 8 +++++--- 1 file changed, 5 insertions(+), 3 deletions(-) diff --git a/src/verilog/Value.v b/src/verilog/Value.v index ad946ca..253595b 100644 --- a/src/verilog/Value.v +++ b/src/verilog/Value.v @@ -381,6 +381,8 @@ Proof. - rewrite wzero'_def. apply wordToNat_wzero. - rewrite posToWord_nat. rewrite wordToNat_natToWord_2. auto. inversion H. destruct (2 ^ Z.of_nat sz)%Z eqn:?; try discriminate. - Set Printing All. - Search positive Z. - - lia. +(* Set Printing All. + Search positive Z.*) + admit. + - admit. +Admitted. -- cgit From 91c19499997ffc7f1e78f7bf0908b43615cf67f1 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Mon, 22 Jun 2020 18:38:22 +0100 Subject: Start Aindexed proof. --- src/translation/HTLgen.v | 9 +- src/translation/HTLgenproof.v | 236 +++++++++++++++++++++++++++++++++++++++--- 2 files changed, 226 insertions(+), 19 deletions(-) diff --git a/src/translation/HTLgen.v b/src/translation/HTLgen.v index 2364a0f..cc6a3f8 100644 --- a/src/translation/HTLgen.v +++ b/src/translation/HTLgen.v @@ -252,7 +252,10 @@ Definition check_address_parameter (p : Z) : bool := Definition translate_eff_addressing (a: Op.addressing) (args: list reg) : mon expr := match a, args with (* TODO: We should be more methodical here; what are the possibilities?*) - | Op.Aindexed off, r1::nil => ret (boplitz Vadd r1 off) + | Op.Aindexed off, r1::nil => + if (check_address_parameter off) + then ret (boplitz Vadd r1 off) + else error (Errors.msg "Veriloggen: translate_eff_addressing address misaligned") | Op.Ascaled scale offset, r1::nil => if (check_address_parameter scale) && (check_address_parameter offset) then ret (Vbinop Vadd (boplitz Vmul r1 scale) (Vlit (ZToValue 32 offset))) @@ -345,7 +348,9 @@ Definition translate_arr_access (mem : AST.memory_chunk) (addr : Op.addressing) (args : list reg) (stack : reg) : mon expr := match mem, addr, args with (* TODO: We should be more methodical here; what are the possibilities?*) | Mint32, Op.Aindexed off, r1::nil => - ret (Vvari stack (Vbinop Vadd (boplitz Vdiv r1 4) (Vlit (ZToValue 32 (off / 4))))) + if (check_address_parameter off) + then ret (Vvari stack (Vbinop Vadd (boplitz Vdiv r1 4) (Vlit (ZToValue 32 (off / 4))))) + else error (Errors.msg "Veriloggen: translate_arr_access address misaligned") | Mint32, Op.Ascaled scale offset, r1::nil => if (check_address_parameter scale) && (check_address_parameter offset) then ret (Vvari stack (Vbinop Vadd (boplitz Vmul r1 (scale / 4)) (Vlit (ZToValue 32 (offset / 4))))) diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index 21dbc73..a3a969c 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -488,17 +488,212 @@ Section CORRECTNESS. | [ _ : context[match ?x with | _ => _ end] |- _ ] => destruct x end. - - (* FIXME: Should be able to use the spec to avoid destructing here. *) + - (* FIXME: Should be able to use the spec to avoid destructing here? *) destruct c, chunk, addr, args; simplify; rt; simplify. - + admit. (* FIXME: Alignment *) + + (** Preamble *) + invert MARR. simplify. + + unfold Op.eval_addressing in H0. + destruct (Archi.ptr64) eqn:ARCHI; simplify. + + unfold reg_stack_based_pointers in RSBP. + pose proof (RSBP r0) as RSBPr0. + + destruct (Registers.Regmap.get r0 rs) eqn:EQr0; simplify. + + rewrite ARCHI in H1. simplify. + subst. + + pose proof MASSOC as MASSOC'. + invert MASSOC'. + pose proof (H0 r0). + assert (HPler0 : Ple r0 (RTL.max_reg_function f)) + by (eapply RTL.max_reg_function_use; eauto; simplify; eauto). + apply H4 in HPler0. + invert HPler0; try congruence. + rewrite EQr0 in H6. + invert H6. + clear H0. clear H4. + + unfold check_address_parameter in *. simplify. + + (** Normalisation proof *) + assert + (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0)) + (Integers.Ptrofs.of_int (Integers.Int.repr z)) + = + Integers.Ptrofs.repr + (4 * + Integers.Ptrofs.unsigned + (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0 / 4)) + (Integers.Ptrofs.of_int (Integers.Int.repr (z / 4)))))) + as NORMALISE. + etransitivity. + 2: { + replace (4 * + Integers.Ptrofs.unsigned + (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0 / 4)) + (Integers.Ptrofs.of_int (Integers.Int.repr (z / 4))))) + with (Integers.Ptrofs.unsigned (Integers.Ptrofs.repr 4) * + Integers.Ptrofs.unsigned + (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0 / 4)) + (Integers.Ptrofs.of_int (Integers.Int.repr (z / 4))))). + 2: { rewrite Integers.Ptrofs.unsigned_repr_eq. + rewrite Z.mod_small. reflexivity. + simplify. lia. } + + rewrite <- PtrofsExtra.mul_unsigned. + rewrite Integers.Ptrofs.mul_add_distr_r. + rewrite PtrofsExtra.mul_repr; simplify; try lia. + rewrite ZLib.div_mul_undo; try lia. + rewrite mul_of_int; simplify; try lia. + rewrite IntExtra.mul_repr; simplify; try lia. + rewrite ZLib.div_mul_undo; try lia. + reflexivity. + + split. + apply Z.div_le_lower_bound; lia. + apply Z.div_le_upper_bound; lia. + + admit. (* FIXME: Register size 32 bits *) } + reflexivity. + + remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0)) + (Integers.Ptrofs.of_int (Integers.Int.repr z))) as OFFSET. + + (** Read bounds assumption *) + assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH by admit. + + eexists. split. + eapply Smallstep.plus_one. + eapply HTL.step_module; eauto. + econstructor. econstructor. econstructor. simplify. + econstructor. econstructor. econstructor. simplify. + eapply Verilog.erun_Vbinop with (EQ := ?[EQ1]). (* FIXME: These will be shelved and cause sadness. *) + eapply Verilog.erun_Vbinop with (EQ := ?[EQ2]). + econstructor. econstructor. econstructor. econstructor. + econstructor. econstructor. econstructor. econstructor. + + (** Verilog array lookup *) + unfold Verilog.arr_assocmap_lookup. simplify. setoid_rewrite H3. + f_equal. + + (** State Lookup *) + unfold Verilog.merge_regs. + simplify. + unfold_merge. + rewrite AssocMap.gso. + apply AssocMap.gss. + apply st_greater_than_res. + + (** Match states *) + rewrite assumption_32bit. + econstructor; eauto. + + (** Match assocmaps *) + unfold Verilog.merge_regs. simplify. unfold_merge. + apply regs_lessdef_add_match. + + (** Equality proof *) + assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) by reflexivity. + rewrite H0 in H5. clear H0. + setoid_rewrite Integers.Ptrofs.add_zero_l in H5. + + specialize (H5 (Integers.Ptrofs.unsigned + (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0 / 4)) + (Integers.Ptrofs.of_int (Integers.Int.repr (z / 4)))))). + + exploit H5; auto; intros. + rewrite Z2Nat.id. + split. + apply Integers.Ptrofs.unsigned_range_2. + admit. + apply Z_div_pos; lia. + clear H5. + + assert (Z.to_nat + (Integers.Ptrofs.unsigned + (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0 / 4)) + (Integers.Ptrofs.of_int (Integers.Int.repr (z / 4))))) + = + valueToNat (vplus (vdivs asr # r0 (ZToValue 32 4) ?EQ2) (ZToValue 32 (z / 4)) ?EQ1)) + as EXPR_OK by admit. + + rewrite <- EXPR_OK. + rewrite <- NORMALISE in H0. + rewrite H1 in H0. + invert H0. assumption. + + (** PC match *) + apply regs_lessdef_add_greater. + apply greater_than_max_func. + assumption. + + (** States well formed *) + unfold state_st_wf. inversion 1. simplify. + unfold Verilog.merge_regs. + unfold_merge. rewrite AssocMap.gso. + apply AssocMap.gss. + apply st_greater_than_res. + + (** Match arrays *) + econstructor. + repeat split; simplify. + unfold HTL.empty_stack. + simplify. + unfold Verilog.merge_arrs. + + rewrite AssocMap.gcombine. + 2: { reflexivity. } + rewrite AssocMap.gss. + unfold Verilog.merge_arr. + setoid_rewrite H3. + reflexivity. + + rewrite combine_length. + unfold arr_repeat. simplify. + rewrite list_repeat_len. + reflexivity. + + unfold arr_repeat. simplify. + rewrite list_repeat_len. + congruence. + + intros. + erewrite array_get_error_equal. + eauto. apply combine_none. + assumption. + + (** RSBP preservation *) + unfold reg_stack_based_pointers. intros. + destruct (Pos.eq_dec r1 dst); subst. + + rewrite Registers.Regmap.gss. + unfold arr_stack_based_pointers in ASBP. + specialize (ASBP (Integers.Ptrofs.unsigned + (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0 / 4)) + (Integers.Ptrofs.of_int (Integers.Int.repr (z / 4)))))). + exploit ASBP; auto; intros. + admit. + + rewrite <- NORMALISE in H0. + assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) by reflexivity. + rewrite H10 in H0. clear H10. + simplify. + rewrite Integers.Ptrofs.add_zero_l in H0. + rewrite H1 in H0. + assumption. + + rewrite Registers.Regmap.gso; auto. + (* FIXME: Why is this degenerate? Should we support this mode? *) unfold Op.eval_addressing in H0. destruct (Archi.ptr64) eqn:ARCHI; simplify. destruct (Registers.Regmap.get r0 rs) eqn:EQr0; discriminate. - + invert MARR. simplify. + + (** Preamble *) + invert MARR. simplify. unfold Op.eval_addressing in H0. destruct (Archi.ptr64) eqn:ARCHI; simplify. @@ -598,19 +793,18 @@ Section CORRECTNESS. (Integers.Int.add (Integers.Int.mul (valueToInt asr # r1) (Integers.Int.repr z)) (Integers.Int.repr z0)))) as OFFSET. - (** Here we are assuming that any stack read will be within the bounds - of the activation record. *) + (** Read bounds assumption *) assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH by admit. - (* Proof begins in earnest here. *) + (** Start of proof proper *) eexists. split. eapply Smallstep.plus_one. eapply HTL.step_module; eauto. econstructor. econstructor. econstructor. simplify. econstructor. econstructor. econstructor. simplify. - eapply Verilog.erun_Vbinop with (EQ := ?[EQ1]). (* FIXME: These will be shelved and cause sadness. *) - eapply Verilog.erun_Vbinop with (EQ := ?[EQ2]). - eapply Verilog.erun_Vbinop with (EQ := ?[EQ3]). + eapply Verilog.erun_Vbinop with (EQ := ?[EQ3]). (* FIXME: These will be shelved and cause sadness. *) + eapply Verilog.erun_Vbinop with (EQ := ?[EQ4]). + eapply Verilog.erun_Vbinop with (EQ := ?[EQ5]). econstructor. econstructor. econstructor. @@ -619,7 +813,7 @@ Section CORRECTNESS. econstructor. econstructor. econstructor. - eapply Verilog.erun_Vbinop with (EQ := ?[EQ4]). + eapply Verilog.erun_Vbinop with (EQ := ?[EQ6]). econstructor. econstructor. econstructor. @@ -628,19 +822,24 @@ Section CORRECTNESS. econstructor. econstructor. (* FIXME: Oh dear. *) + (** Verilog array lookup *) unfold Verilog.arr_assocmap_lookup. simplify. setoid_rewrite H3. f_equal. - simplify. unfold Verilog.merge_regs. + (** State Lookup *) + unfold Verilog.merge_regs. + simplify. unfold_merge. rewrite AssocMap.gso. apply AssocMap.gss. apply st_greater_than_res. - simplify. rewrite assumption_32bit. - econstructor; simplify; eauto. + (** Match states *) + rewrite assumption_32bit. + econstructor; eauto. - unfold Verilog.merge_regs. unfold_merge. + (** Match assocmaps *) + unfold Verilog.merge_regs. simplify. unfold_merge. apply regs_lessdef_add_match. (** Equality proof *) @@ -674,23 +873,26 @@ Section CORRECTNESS. (Integers.Int.repr (z0 / 4)))))) = valueToNat - (vplus (vplus (vdivs asr # r0 (ZToValue 32 4) ?EQ3) (ZToValue 32 (z0 / 4)) ?EQ2) - (vmul asr # r1 (ZToValue 32 (z / 4)) ?EQ4) ?EQ1)) as EXPR_OK by admit. + (vplus (vplus (vdivs asr # r0 (ZToValue 32 4) ?EQ5) (ZToValue 32 (z0 / 4)) ?EQ4) + (vmul asr # r1 (ZToValue 32 (z / 4)) ?EQ6) ?EQ3)) as EXPR_OK by admit. rewrite <- EXPR_OK. rewrite <- NORMALISE in H0. rewrite H1 in H0. invert H0. assumption. + (** PC match *) apply regs_lessdef_add_greater. apply greater_than_max_func. assumption. + (** States well formed *) unfold state_st_wf. inversion 1. simplify. unfold Verilog.merge_regs. unfold_merge. rewrite AssocMap.gso. apply AssocMap.gss. apply st_greater_than_res. + (** Match arrays *) econstructor. repeat split; simplify. unfold HTL.empty_stack. @@ -718,6 +920,7 @@ Section CORRECTNESS. eauto. apply combine_none. assumption. + (** RSBP preservation *) unfold reg_stack_based_pointers. intros. destruct (Pos.eq_dec r2 dst); subst. @@ -862,7 +1065,6 @@ Section CORRECTNESS. destruct (Pos.eq_dec r0 dst); subst. rewrite Registers.Regmap.gss. - unfold arr_stack_based_pointers in ASBP. specialize (ASBP (Integers.Ptrofs.unsigned i0 / 4)). exploit ASBP; auto; intros. -- cgit From b70f007eae5886990a8ffc1e7b94294702e238f8 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Mon, 22 Jun 2020 18:46:26 +0100 Subject: Finish off Load proof sketches. --- src/translation/HTLgenproof.v | 52 ++++++++++++++----------------------------- 1 file changed, 17 insertions(+), 35 deletions(-) diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index a3a969c..7e9cb26 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -684,6 +684,7 @@ Section CORRECTNESS. rewrite Integers.Ptrofs.add_zero_l in H0. rewrite H1 in H0. assumption. + simplify. rewrite Registers.Regmap.gso; auto. @@ -925,44 +926,25 @@ Section CORRECTNESS. destruct (Pos.eq_dec r2 dst); subst. rewrite Registers.Regmap.gss. - unfold arr_stack_based_pointers in ASBP. - (* specialize (ASBP ((Integers.Ptrofs.unsigned *) - (* (Integers.Ptrofs.add i0 *) - (* (Integers.Ptrofs.of_int *) - (* (Integers.Int.add (Integers.Int.mul i1 (Integers.Int.repr z)) *) - (* (Integers.Int.repr z0))))) / 4)). *) - (* exploit ASBP; auto; intros. *) - (* 1: { *) - (* split. *) - (* - admit. (*apply Z.div_pos; lia.*) *) - (* - admit. (* apply Zmult_lt_reg_r with (p := 4); try lia. *) *) - (* (* repeat rewrite ZLib.div_mul_undo; lia. *) *) - (* } *) - (* replace (4 * ((Integers.Ptrofs.unsigned *) - (* (Integers.Ptrofs.add i0 *) - (* (Integers.Ptrofs.of_int *) - (* (Integers.Int.add (Integers.Int.mul i1 (Integers.Int.repr z)) *) - (* (Integers.Int.repr z0))))) / 4)) with (Integers.Ptrofs.unsigned *) - (* (Integers.Ptrofs.add i0 *) - (* (Integers.Ptrofs.of_int *) - (* (Integers.Int.add (Integers.Int.mul i1 (Integers.Int.repr z)) *) - (* (Integers.Int.repr z0))))) in H0. *) - (* 2: { *) - (* rewrite Z.mul_comm. *) - (* admit. *) - (* (* rewrite ZLib.div_mul_undo; lia. *) *) - (* } *) - (* rewrite Integers.Ptrofs.repr_unsigned in H0. *) - (* simplify. *) - (* assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) by reflexivity. *) - (* rewrite H4 in H0. *) - (* setoid_rewrite Integers.Ptrofs.add_zero_l in H0. *) - (* rewrite H1 in H0. *) - (* simplify. *) - (* assumption. *) + specialize (ASBP (Integers.Ptrofs.unsigned + (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0 / 4)) + (Integers.Ptrofs.of_int + (Integers.Int.add + (Integers.Int.mul (valueToInt asr # r1) (Integers.Int.repr (z / 4))) + (Integers.Int.repr (z0 / 4))))))). + exploit ASBP; auto; intros. admit. + rewrite <- NORMALISE in H0. + assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) by reflexivity. + rewrite H19 in H0. clear H19. + simplify. + rewrite Integers.Ptrofs.add_zero_l in H0. + rewrite H1 in H0. + assumption. + simplify. + rewrite Registers.Regmap.gso; auto. + invert MARR. simplify. -- cgit From 8c48e9e1094f037835698c88782772c8b3250a76 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Tue, 23 Jun 2020 11:34:16 +0100 Subject: More to proof --- src/verilog/Value.v | 16 +++++++++------- 1 file changed, 9 insertions(+), 7 deletions(-) diff --git a/src/verilog/Value.v b/src/verilog/Value.v index 253595b..bde98b8 100644 --- a/src/verilog/Value.v +++ b/src/verilog/Value.v @@ -374,15 +374,17 @@ Proof. destruct b; unfold valueToBool, boolToValue; simpl; trivial. Qed. Lemma ZToValue_valueToNat : forall x sz, + sz > 0 -> (x < 2^(Z.of_nat sz))%Z -> valueToNat (ZToValue sz x) = Z.to_nat x. Proof. destruct x; intros; unfold ZToValue, valueToNat; simpl. - rewrite wzero'_def. apply wordToNat_wzero. - - rewrite posToWord_nat. rewrite wordToNat_natToWord_2. auto. - inversion H. destruct (2 ^ Z.of_nat sz)%Z eqn:?; try discriminate. -(* Set Printing All. - Search positive Z.*) - admit. - - admit. -Admitted. + - rewrite posToWord_nat. rewrite wordToNat_natToWord_2. trivial. + unfold Z.of_nat in *. destruct sz eqn:?. omega. simpl in H0. + rewrite <- Pos2Z.inj_pow_pos in H0. Search (Z.pos _ < Z.pos _)%Z. + Search Pos.to_nat (_ < _). (* Pos2Nat.inj_lt *) + Search Pos.to_nat. + (* Pos2Nat.is_succ: forall p : positive, exists n : nat, Pos.to_nat p = S n *) + Search S Pos.to_nat. Search (_ < _ <-> _)%positive. assert (exists p, Pos.to_nat p = S n). + econstructor. Search Pos2Nat.is_succ. Search Pos.of_succ_nat. -- cgit From ec0fa1ac249a8eeb0df9700c50a3e6c4f1b540f2 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Tue, 23 Jun 2020 23:00:08 +0100 Subject: Normalise entire expression to avoid overflow issues. --- src/common/Coquplib.v | 10 + src/common/IntegerExtra.v | 235 +++++++++++++++++++ src/translation/HTLgen.v | 11 +- src/translation/HTLgenproof.v | 532 +++++++++++++++++++++--------------------- src/translation/HTLgenspec.v | 7 +- 5 files changed, 523 insertions(+), 272 deletions(-) create mode 100644 src/common/IntegerExtra.v diff --git a/src/common/Coquplib.v b/src/common/Coquplib.v index 59b23ae..b4ca906 100644 --- a/src/common/Coquplib.v +++ b/src/common/Coquplib.v @@ -84,6 +84,11 @@ Ltac unfold_constants := | [ H : context[Integers.Ptrofs.max_signed] |- _ ] => replace Integers.Ptrofs.max_signed with 2147483647 in H by reflexivity + | [ _ : _ |- context[Integers.Ptrofs.max_unsigned] ] => + replace Integers.Ptrofs.max_unsigned with 4294967295 by reflexivity + | [ H : context[Integers.Ptrofs.max_unsigned] |- _ ] => + replace Integers.Ptrofs.max_unsigned with 4294967295 in H by reflexivity + | [ _ : _ |- context[Integers.Int.modulus] ] => replace Integers.Int.modulus with 4294967296 by reflexivity | [ H : context[Integers.Int.modulus] |- _ ] => @@ -98,6 +103,11 @@ Ltac unfold_constants := replace Integers.Int.max_signed with 2147483647 by reflexivity | [ H : context[Integers.Int.max_signed] |- _ ] => replace Integers.Int.max_signed with 2147483647 in H by reflexivity + + | [ _ : _ |- context[Integers.Int.max_unsigned] ] => + replace Integers.Int.max_unsigned with 4294967295 by reflexivity + | [ H : context[Integers.Int.max_unsigned] |- _ ] => + replace Integers.Int.max_unsigned with 4294967295 in H by reflexivity end. Ltac simplify := unfold_constants; simpl in *; diff --git a/src/common/IntegerExtra.v b/src/common/IntegerExtra.v new file mode 100644 index 0000000..ad01015 --- /dev/null +++ b/src/common/IntegerExtra.v @@ -0,0 +1,235 @@ +Require Import BinInt. +Require Import Lia. +Require Import ZBinary. + +From bbv Require Import ZLib. +From compcert Require Import Integers Coqlib. + +Require Import Coquplib. + +Local Open Scope Z_scope. + +Module PtrofsExtra. + + Lemma mul_divs : + forall x y, + 0 <= Ptrofs.signed y -> + 0 < Ptrofs.signed x -> + Ptrofs.signed y mod Ptrofs.signed x = 0 -> + (Integers.Ptrofs.mul x (Integers.Ptrofs.divs y x)) = y. + Proof. + intros. + + pose proof (Ptrofs.mods_divs_Euclid y x). + pose proof (Zquot.Zrem_Zmod_zero (Ptrofs.signed y) (Ptrofs.signed x)). + apply <- H3 in H1; try lia; clear H3. + unfold Ptrofs.mods in H2. + rewrite H1 in H2. + replace (Ptrofs.repr 0) with (Ptrofs.zero) in H2 by reflexivity. + rewrite Ptrofs.add_zero in H2. + rewrite Ptrofs.mul_commut. + congruence. + Qed. + + Lemma Z_div_distr_add : + forall x y z, + x mod z = 0 -> + y mod z = 0 -> + z <> 0 -> + x / z + y / z = (x + y) / z. + Proof. + intros. + + assert ((x + y) mod z = 0). + { rewrite <- Z.add_mod_idemp_l; try assumption. + rewrite H. assumption. } + + rewrite <- Z.mul_cancel_r with (p := z); try assumption. + rewrite Z.mul_add_distr_r. + repeat rewrite ZLib.div_mul_undo; lia. + Qed. + + Lemma mul_unsigned : + forall x y, + Ptrofs.mul x y = + Ptrofs.repr (Ptrofs.unsigned x * Ptrofs.unsigned y). + Proof. + intros; unfold Ptrofs.mul. + apply Ptrofs.eqm_samerepr. + apply Ptrofs.eqm_mult; exists 0; lia. + Qed. + + Lemma mul_repr : + forall x y, + Ptrofs.min_signed <= y <= Ptrofs.max_signed -> + Ptrofs.min_signed <= x <= Ptrofs.max_signed -> + Ptrofs.mul (Ptrofs.repr y) (Ptrofs.repr x) = Ptrofs.repr (x * y). + Proof. + intros; unfold Ptrofs.mul. + destruct (Z_ge_lt_dec x 0); destruct (Z_ge_lt_dec y 0). + + - f_equal. + repeat rewrite Ptrofs.unsigned_repr_eq. + repeat rewrite Z.mod_small; simplify; lia. + + - assert (Ptrofs.lt (Ptrofs.repr y) Ptrofs.zero = true). + { + unfold Ptrofs.lt. + rewrite Ptrofs.signed_repr; auto. + rewrite Ptrofs.signed_zero. + destruct (zlt y 0); try lia; auto. + } + + rewrite Ptrofs.unsigned_signed with (n := Ptrofs.repr y). + rewrite H1. + rewrite Ptrofs.signed_repr; auto. + rewrite Ptrofs.unsigned_repr_eq. + rewrite Z.mod_small; simplify; try lia. + rewrite Z.mul_add_distr_r. + apply Ptrofs.eqm_samerepr. + exists x. simplify. lia. + + - assert (Ptrofs.lt (Ptrofs.repr x) Ptrofs.zero = true). + { + unfold Ptrofs.lt. + rewrite Ptrofs.signed_repr; auto. + rewrite Ptrofs.signed_zero. + destruct (zlt x 0); try lia; auto. + } + + rewrite Ptrofs.unsigned_signed with (n := Ptrofs.repr x). + rewrite H1. + rewrite Ptrofs.signed_repr; auto. + rewrite Ptrofs.unsigned_repr_eq. + rewrite Z.mod_small; simplify; try lia. + rewrite Z.mul_add_distr_l. + apply Ptrofs.eqm_samerepr. + exists y. simplify. lia. + + - assert (Ptrofs.lt (Ptrofs.repr x) Ptrofs.zero = true). + { + unfold Ptrofs.lt. + rewrite Ptrofs.signed_repr; auto. + rewrite Ptrofs.signed_zero. + destruct (zlt x 0); try lia; auto. + } + assert (Ptrofs.lt (Ptrofs.repr y) Ptrofs.zero = true). + { + unfold Ptrofs.lt. + rewrite Ptrofs.signed_repr; auto. + rewrite Ptrofs.signed_zero. + destruct (zlt y 0); try lia; auto. + } + rewrite Ptrofs.unsigned_signed with (n := Ptrofs.repr x). + rewrite Ptrofs.unsigned_signed with (n := Ptrofs.repr y). + rewrite H2. + rewrite H1. + repeat rewrite Ptrofs.signed_repr; auto. + replace ((y + Ptrofs.modulus) * (x + Ptrofs.modulus)) with + (x * y + (x + y + Ptrofs.modulus) * Ptrofs.modulus) by lia. + apply Ptrofs.eqm_samerepr. + exists (x + y + Ptrofs.modulus). lia. + Qed. +End PtrofsExtra. + +Module IntExtra. + Lemma mul_unsigned : + forall x y, + Int.mul x y = + Int.repr (Int.unsigned x * Int.unsigned y). + Proof. + intros; unfold Int.mul. + apply Int.eqm_samerepr. + apply Int.eqm_mult; exists 0; lia. + Qed. + + Lemma mul_repr : + forall x y, + Int.min_signed <= y <= Int.max_signed -> + Int.min_signed <= x <= Int.max_signed -> + Int.mul (Int.repr y) (Int.repr x) = Int.repr (x * y). + Proof. + intros; unfold Int.mul. + destruct (Z_ge_lt_dec x 0); destruct (Z_ge_lt_dec y 0). + + - f_equal. + repeat rewrite Int.unsigned_repr_eq. + repeat rewrite Z.mod_small; simplify; lia. + + - assert (Int.lt (Int.repr y) Int.zero = true). + { + unfold Int.lt. + rewrite Int.signed_repr; auto. + rewrite Int.signed_zero. + destruct (zlt y 0); try lia; auto. + } + + rewrite Int.unsigned_signed with (n := Int.repr y). + rewrite H1. + rewrite Int.signed_repr; auto. + rewrite Int.unsigned_repr_eq. + rewrite Z.mod_small; simplify; try lia. + rewrite Z.mul_add_distr_r. + apply Int.eqm_samerepr. + exists x. simplify. lia. + + - assert (Int.lt (Int.repr x) Int.zero = true). + { + unfold Int.lt. + rewrite Int.signed_repr; auto. + rewrite Int.signed_zero. + destruct (zlt x 0); try lia; auto. + } + + rewrite Int.unsigned_signed with (n := Int.repr x). + rewrite H1. + rewrite Int.signed_repr; auto. + rewrite Int.unsigned_repr_eq. + rewrite Z.mod_small; simplify; try lia. + rewrite Z.mul_add_distr_l. + apply Int.eqm_samerepr. + exists y. simplify. lia. + + - assert (Int.lt (Int.repr x) Int.zero = true). + { + unfold Int.lt. + rewrite Int.signed_repr; auto. + rewrite Int.signed_zero. + destruct (zlt x 0); try lia; auto. + } + assert (Int.lt (Int.repr y) Int.zero = true). + { + unfold Int.lt. + rewrite Int.signed_repr; auto. + rewrite Int.signed_zero. + destruct (zlt y 0); try lia; auto. + } + rewrite Int.unsigned_signed with (n := Int.repr x). + rewrite Int.unsigned_signed with (n := Int.repr y). + rewrite H2. + rewrite H1. + repeat rewrite Int.signed_repr; auto. + replace ((y + Int.modulus) * (x + Int.modulus)) with + (x * y + (x + y + Int.modulus) * Int.modulus) by lia. + apply Int.eqm_samerepr. + exists (x + y + Int.modulus). lia. + Qed. +End IntExtra. + +Lemma mul_of_int : + forall x y, + 0 <= x < Integers.Ptrofs.modulus -> + Integers.Ptrofs.mul (Integers.Ptrofs.repr x) (Integers.Ptrofs.of_int y) = + Integers.Ptrofs.of_int (Integers.Int.mul (Integers.Int.repr x) y). +Proof. + intros. + pose proof (Integers.Ptrofs.agree32_of_int eq_refl y) as A. + pose proof (Integers.Ptrofs.agree32_to_int eq_refl (Integers.Ptrofs.repr x)) as B. + exploit Integers.Ptrofs.agree32_mul; [> reflexivity | exact B | exact A | intro C]. + unfold Integers.Ptrofs.to_int in C. + unfold Integers.Ptrofs.of_int in C. + rewrite Integers.Ptrofs.unsigned_repr_eq in C. + rewrite Z.mod_small in C; auto. + symmetry. + apply Integers.Ptrofs.agree32_of_int_eq; auto. +Qed. diff --git a/src/translation/HTLgen.v b/src/translation/HTLgen.v index cc6a3f8..92e40f5 100644 --- a/src/translation/HTLgen.v +++ b/src/translation/HTLgen.v @@ -349,17 +349,18 @@ Definition translate_arr_access (mem : AST.memory_chunk) (addr : Op.addressing) match mem, addr, args with (* TODO: We should be more methodical here; what are the possibilities?*) | Mint32, Op.Aindexed off, r1::nil => if (check_address_parameter off) - then ret (Vvari stack (Vbinop Vadd (boplitz Vdiv r1 4) (Vlit (ZToValue 32 (off / 4))))) + then ret (Vvari stack (Vbinop Vdiv (boplitz Vadd r1 off) (Vlit (ZToValue 32 4)))) else error (Errors.msg "Veriloggen: translate_arr_access address misaligned") | Mint32, Op.Ascaled scale offset, r1::nil => if (check_address_parameter scale) && (check_address_parameter offset) - then ret (Vvari stack (Vbinop Vadd (boplitz Vmul r1 (scale / 4)) (Vlit (ZToValue 32 (offset / 4))))) + then ret (Vvari stack (Vbinop Vdiv (Vbinop Vadd (boplitz Vmul r1 scale) (Vlit (ZToValue 32 offset))) (Vlit (ZToValue 32 4)))) else error (Errors.msg "Veriloggen: translate_arr_access address misaligned") | Mint32, Op.Aindexed2scaled scale offset, r1::r2::nil => (* Typical for dynamic array addressing *) if (check_address_parameter scale) && (check_address_parameter offset) then ret (Vvari stack - (Vbinop Vadd (Vbinop Vadd (boplitz Vdiv r1 4) (Vlit (ZToValue 32 (offset / 4)))) - (boplitz Vmul r2 (scale / 4)))) + (Vbinop Vdiv + (Vbinop Vadd (boplitz Vadd r1 offset) (boplitz Vmul r2 scale)) + (ZToValue 32 4))) else error (Errors.msg "Veriloggen: translate_arr_access address misaligned") | Mint32, Op.Ainstack a, nil => (* We need to be sure that the base address is aligned *) let a := Integers.Ptrofs.unsigned a in @@ -451,7 +452,7 @@ Definition create_arr (i : option io) (sz : nat) (ln : nat) : mon (reg * nat) := (create_arr_state_incr s sz ln i). Definition stack_correct (sz : Z) : bool := - (0 <=? sz) && (Z.modulo sz 4 =? 0). + (0 <=? sz) && (sz Prop := tr_code f.(RTL.fn_code) pc i data control fin rtrn st stk) -> stk_len = Z.to_nat (f.(RTL.fn_stacksize) / 4) -> Z.modulo (f.(RTL.fn_stacksize)) 4 = 0 -> - 0 <= f.(RTL.fn_stacksize) -> + 0 <= f.(RTL.fn_stacksize) < Integers.Ptrofs.modulus -> m = (mkmodule f.(RTL.fn_params) data control @@ -454,7 +454,8 @@ Proof. unfold transf_module in *. unfold stack_correct in *. - destruct (0 <=? RTL.fn_stacksize f) eqn:STACK_BOUND; + destruct (0 <=? RTL.fn_stacksize f) eqn:STACK_BOUND_LOW; + destruct (RTL.fn_stacksize f Date: Wed, 24 Jun 2020 10:45:16 +0100 Subject: Fix assumption of main --- src/translation/HTLgen.v | 22 +++++++----- src/translation/HTLgenproof.v | 77 ++++++++++++++++++++++++++++++---------- src/translation/HTLgenspec.v | 81 +++++++++++++------------------------------ 3 files changed, 97 insertions(+), 83 deletions(-) diff --git a/src/translation/HTLgen.v b/src/translation/HTLgen.v index f109a8e..0482c61 100644 --- a/src/translation/HTLgen.v +++ b/src/translation/HTLgen.v @@ -283,7 +283,6 @@ Definition translate_condition (c : Op.condition) (args : list reg) : mon expr : Definition translate_eff_addressing (a: Op.addressing) (args: list reg) : mon expr := match a, args with | Op.Aindexed off, r1::nil => ret (boplitz Vadd r1 off) - | Op.Aindexed2 off, r1::r2::nil => ret (Vbinop Vadd (Vvar r1) (boplitz Vadd r2 off)) | Op.Ascaled scale offset, r1::nil => ret (Vbinop Vadd (boplitz Vmul r1 scale) (Vlit (ZToValue 32 offset))) | Op.Aindexed2scaled scale offset, r1::r2::nil => @@ -529,7 +528,7 @@ Definition transl_module (f : function) : Errors.res module := Definition transl_fundef := transf_partial_fundef transl_module. -Definition transl_program (p : RTL.program) := transform_partial_program transl_fundef p. +(* Definition transl_program (p : RTL.program) := transform_partial_program transl_fundef p. *) (*Definition transl_main_fundef f : Errors.res HTL.fundef := match f with @@ -546,11 +545,18 @@ Definition transl_program (p: RTL.program) : Errors.res HTL.program := (fun i v => Errors.OK v) p. *) -(*Definition main_is_internal (p : RTL.program): Prop := +Definition main_is_internal {A B : Type} (p : AST.program A B) : bool := let ge := Globalenvs.Genv.globalenv p in - forall b m, - Globalenvs.Genv.find_symbol ge p.(AST.prog_main) = Some b -> - Globalenvs.Genv.find_funct_ptr ge b = Some (AST.Internal m). + match Globalenvs.Genv.find_symbol ge p.(AST.prog_main) with + | Some b => + match Globalenvs.Genv.find_funct_ptr ge b with + | Some (AST.Internal _) => true + | _ => false + end + | _ => false + end. -Definition tranls_program_with_proof (p : RTL.program) : Errors.res { p' : HTL.program | main_is_internal p }. -*) +Definition transl_program (p : RTL.program) : Errors.res HTL.program := + if main_is_internal p + then transform_partial_program transl_fundef p + else Errors.Error (Errors.msg "Main function is not Internal."). diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index 9786d23..c40b737 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -87,12 +87,23 @@ Inductive match_states : RTL.state -> HTL.state -> Prop := Hint Constructors match_states : htlproof. Definition match_prog (p: RTL.program) (tp: HTL.program) := - Linking.match_program (fun cu f tf => transl_fundef f = Errors.OK tf) eq p tp. + Linking.match_program (fun cu f tf => transl_fundef f = Errors.OK tf) eq p tp /\ + main_is_internal p = true. + +Definition match_prog_matches : + forall p tp, + match_prog p tp -> + Linking.match_program (fun cu f tf => transl_fundef f = Errors.OK tf) eq p tp. + Proof. intros. unfold match_prog in H. tauto. Qed. Lemma transf_program_match: forall p tp, HTLgen.transl_program p = Errors.OK tp -> match_prog p tp. Proof. - intros. apply Linking.match_transform_partial_program; auto. + intros. unfold transl_program in H. + destruct (main_is_internal p) eqn:?; try discriminate. + unfold match_prog. split. + apply Linking.match_transform_partial_program; auto. + assumption. Qed. Lemma regs_lessdef_add_greater : @@ -205,13 +216,16 @@ Section CORRECTNESS. Hypothesis TRANSL : match_prog prog tprog. -(* Let ge : RTL.genv := Globalenvs.Genv.globalenv prog. + Lemma TRANSL' : + Linking.match_program (fun cu f tf => transl_fundef f = Errors.OK tf) eq prog tprog. + Proof. intros; apply match_prog_matches; assumption. Qed. + + Let ge : RTL.genv := Globalenvs.Genv.globalenv prog. Let tge : HTL.genv := Globalenvs.Genv.globalenv tprog. Lemma symbols_preserved: forall (s: AST.ident), Genv.find_symbol tge s = Genv.find_symbol ge s. - Proof - (Genv.find_symbol_transf_partial TRANSL). + Proof. intros. eapply (Genv.find_symbol_match TRANSL'). Qed. Lemma function_ptr_translated: forall (b: Values.block) (f: RTL.fundef), @@ -219,7 +233,7 @@ Section CORRECTNESS. exists tf, Genv.find_funct_ptr tge b = Some tf /\ transl_fundef f = Errors.OK tf. Proof. - intros. exploit (Genv.find_funct_ptr_match TRANSL); eauto. + intros. exploit (Genv.find_funct_ptr_match TRANSL'); eauto. intros (cu & tf & P & Q & R); exists tf; auto. Qed. @@ -229,21 +243,21 @@ Section CORRECTNESS. exists tf, Genv.find_funct tge v = Some tf /\ transl_fundef f = Errors.OK tf. Proof. - intros. exploit (Genv.find_funct_match TRANSL); eauto. + intros. exploit (Genv.find_funct_match TRANSL'); eauto. intros (cu & tf & P & Q & R); exists tf; auto. Qed. Lemma senv_preserved: Senv.equiv (Genv.to_senv ge) (Genv.to_senv tge). Proof - (Genv.senv_transf_partial TRANSL). + (Genv.senv_transf_partial TRANSL'). Hint Resolve senv_preserved : htlproof. Lemma eval_correct : - forall sp op rs args m v v' e asr asa f, + forall sp op rs args m v v' e asr asa f s s' i, 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 -> + translate_instr op args s = OK e s' i -> val_value_lessdef v v' -> Verilog.expr_runp f asr asa e v'. Admitted. @@ -441,6 +455,7 @@ Section CORRECTNESS. assumption. + - admit. - (* Return *) econstructor. split. eapply Smallstep.plus_two. @@ -540,6 +555,17 @@ Section CORRECTNESS. Admitted. Hint Resolve transl_step_correct : htlproof. + Lemma main_tprog_internal : + forall b f, + Globalenvs.Genv.find_symbol tge tprog.(AST.prog_main) = Some b -> + Genv.find_funct_ptr (Genv.globalenv tprog) b = Some (AST.Internal f). + Admitted. + + Lemma option_inv : + forall A x y, + @Some A x = Some y -> x = y. + Proof. intros. inversion H. trivial. Qed. + (* Had to admit proof because currently there is no way to force main to be Internal. *) Lemma transl_initial_states : forall s1 : Smallstep.state (RTL.semantics prog), @@ -548,20 +574,33 @@ Section CORRECTNESS. Smallstep.initial_state (HTL.semantics tprog) s2 /\ match_states s1 s2. Proof. induction 1. + destruct TRANSL. unfold main_is_internal in H4. + repeat (unfold_match H4). + assert (f = AST.Internal f1). apply option_inv. + rewrite <- Heqo0. rewrite <- H1. replace b with b0. + auto. apply option_inv. rewrite <- H0. rewrite <- Heqo. + trivial. exploit function_ptr_translated; eauto. intros [tf [A B]]. unfold transl_fundef, Errors.bind in B. + unfold AST.transf_partial_fundef, Errors.bind in B. repeat (unfold_match B). inversion B. subst. econstructor; split. econstructor. - apply (Genv.init_mem_transf_partial TRANSL); eauto. + apply (Genv.init_mem_transf_partial TRANSL'); eauto. replace (AST.prog_main tprog) with (AST.prog_main prog). rewrite symbols_preserved; eauto. symmetry; eapply Linking.match_program_main; eauto. - Admitted. - (*eexact A. trivial. + apply main_tprog_internal. replace ge0 with ge in * by auto. + replace b0 with b in *. rewrite symbols_preserved. + replace (AST.prog_main tprog) with (AST.prog_main prog). + assumption. + symmetry; eapply Linking.match_program_main; eauto. + apply option_inv. rewrite <- H0. rewrite <- Heqo. + trivial. + constructor. - apply transl_module_correct. auto. - Qed.*) + apply transl_module_correct. eassumption. + Qed. Hint Resolve transl_initial_states : htlproof. Lemma transl_final_states : @@ -574,16 +613,16 @@ Section CORRECTNESS. Proof. intros. inv H0. inv H. inv H4. inv MS. constructor. trivial. Qed. - Hint Resolve transl_final_states : htlproof.*) + Hint Resolve transl_final_states : htlproof. Theorem transf_program_correct: Smallstep.forward_simulation (RTL.semantics prog) (HTL.semantics tprog). Proof. -(* eapply Smallstep.forward_simulation_plus. + eapply Smallstep.forward_simulation_plus. apply senv_preserved. eexact transl_initial_states. eexact transl_final_states. - exact transl_step_correct.*) -Admitted. + exact transl_step_correct. +Qed. End CORRECTNESS. diff --git a/src/translation/HTLgenspec.v b/src/translation/HTLgenspec.v index 887a696..d0d16ba 100644 --- a/src/translation/HTLgenspec.v +++ b/src/translation/HTLgenspec.v @@ -113,41 +113,13 @@ Ltac monadInv H := statemachine that is created by the translation contains the correct translations for each of the elements *) -Inductive tr_op : Op.operation -> list reg -> expr -> Prop := -| tr_op_Omove : forall r, tr_op Op.Omove (r::nil) (Vvar r) -| tr_op_Ointconst : forall n l, tr_op (Op.Ointconst n) l (Vlit (intToValue n)) -| tr_op_Oneg : forall r, tr_op Op.Oneg (r::nil) (Vunop Vneg (Vvar r)) -| tr_op_Osub : forall r1 r2, tr_op Op.Osub (r1::r2::nil) (bop Vsub r1 r2) -| tr_op_Omul : forall r1 r2, tr_op Op.Omul (r1::r2::nil) (bop Vmul r1 r2) -| tr_op_Omulimm : forall r n, tr_op (Op.Omulimm n) (r::nil) (boplit Vmul r n) -| tr_op_Odiv : forall r1 r2, tr_op Op.Odiv (r1::r2::nil) (bop Vdiv r1 r2) -| tr_op_Odivu : forall r1 r2, tr_op Op.Odivu (r1::r2::nil) (bop Vdivu r1 r2) -| tr_op_Omod : forall r1 r2, tr_op Op.Omod (r1::r2::nil) (bop Vmod r1 r2) -| tr_op_Omodu : forall r1 r2, tr_op Op.Omodu (r1::r2::nil) (bop Vmodu r1 r2) -| tr_op_Oand : forall r1 r2, tr_op Op.Oand (r1::r2::nil) (bop Vand r1 r2) -| tr_op_Oandimm : forall n r, tr_op (Op.Oandimm n) (r::nil) (boplit Vand r n) -| tr_op_Oor : forall r1 r2, tr_op Op.Oor (r1::r2::nil) (bop Vor r1 r2) -| tr_op_Oorimm : forall n r, tr_op (Op.Oorimm n) (r::nil) (boplit Vor r n) -| tr_op_Oxor : forall r1 r2, tr_op Op.Oxor (r1::r2::nil) (bop Vxor r1 r2) -| tr_op_Oxorimm : forall n r, tr_op (Op.Oxorimm n) (r::nil) (boplit Vxor r n) -| tr_op_Onot : forall r, tr_op Op.Onot (r::nil) (Vunop Vnot (Vvar r)) -| tr_op_Oshl : forall r1 r2, tr_op Op.Oshl (r1::r2::nil) (bop Vshl r1 r2) -| tr_op_Oshlimm : forall n r, tr_op (Op.Oshlimm n) (r::nil) (boplit Vshl r n) -| tr_op_Oshr : forall r1 r2, tr_op Op.Oshr (r1::r2::nil) (bop Vshr r1 r2) -| tr_op_Oshrimm : forall n r, tr_op (Op.Oshrimm n) (r::nil) (boplit Vshr r n) -| tr_op_Ocmp : forall c l e s s' i, translate_condition c l s = OK e s' i -> tr_op (Op.Ocmp c) l e -| tr_op_Olea : forall a l e s s' i, translate_eff_addressing a l s = OK e s' i -> tr_op (Op.Olea a) l e -| tr_op_Oleal : forall a l e s s' i, translate_eff_addressing a l s = OK e s' i -> tr_op (Op.Oleal a) l e -| tr_op_Ocast32signed : forall r, tr_op (Op.Ocast32signed) (r::nil) (Vvar r). -Hint Constructors tr_op : htlspec. - Inductive tr_instr (fin rtrn st stk : reg) : RTL.instruction -> stmnt -> stmnt -> Prop := | tr_instr_Inop : forall n, tr_instr fin rtrn st stk (RTL.Inop n) Vskip (state_goto st n) | tr_instr_Iop : - forall n op args e dst, - tr_op op args e -> + forall n op args dst s s' e i, + translate_instr op args s = OK e s' i -> tr_instr fin rtrn st stk (RTL.Iop op args dst n) (Vnonblock (Vvar dst) e) (state_goto st n) | tr_instr_Icond : forall n1 n2 cond args s s' i c, @@ -168,7 +140,11 @@ Inductive tr_instr (fin rtrn st stk : reg) : RTL.instruction -> stmnt -> stmnt - forall mem addr args s s' i c src n, translate_arr_access mem addr args stk s = OK c s' i -> tr_instr fin rtrn st stk (RTL.Istore mem addr args src n) (Vblock c (Vvar src)) - (state_goto st n). + (state_goto st n) +| tr_instr_Ijumptable : + forall cexpr tbl r, + cexpr = tbl_to_case_expr st tbl -> + tr_instr fin rtrn st stk (RTL.Ijumptable r tbl) (Vskip) (Vcase (Vvar r) cexpr (Some Vskip)). Hint Constructors tr_instr : htlspec. Inductive tr_code (c : RTL.code) (pc : RTL.node) (i : RTL.instruction) (stmnts trans : PTree.t stmnt) @@ -315,21 +291,6 @@ Ltac rewrite_states := remember (?x ?s) as c1; remember (?x ?s') as c2; try subst end. -Lemma translate_instr_tr_op : - forall op args s s' e i, - translate_instr op args s = OK e s' i -> - tr_op op args e. -Proof. -(* intros. - destruct op eqn:?; eauto with htlspec; try discriminate; simpl in *; - try (match goal with - [ H: match ?args with _ => _ end _ = _ _ _ |- _ ] => - repeat (destruct args; try discriminate) - end); - monadInv H; constructor. -Qed.*) Admitted. (* FIXME: Currently admitted because added Osel *) -Hint Resolve translate_instr_tr_op : htlspec. - Ltac unfold_match H := match type of H with | context[match ?g with _ => _ end] => destruct g eqn:?; try discriminate @@ -343,7 +304,7 @@ Ltac inv_add_instr' H := end; repeat unfold_match H; inversion H. Ltac inv_add_instr := - match goal with + lazymatch goal with | H: context[add_instr_skip _ _ _] |- _ => inv_add_instr' H | H: context[add_instr_skip _ _] |- _ => @@ -356,6 +317,10 @@ Ltac inv_add_instr := inv_add_instr' H | H: context[add_branch_instr _ _ _ _] |- _ => monadInv H; inv_incr; inv_add_instr + | H: context[add_node_skip _ _ _] |- _ => + inv_add_instr' H + | H: context[add_node_skip _ _] |- _ => + monadInv H; inv_incr; inv_add_instr end. Ltac destruct_optional := @@ -370,7 +335,7 @@ Lemma iter_expand_instr_spec : c!pc = Some instr -> tr_code c pc instr s'.(st_datapath) s'.(st_controllogic) fin rtrn s'.(st_st) stack). Proof. -(* induction l; simpl; intros; try contradiction. + induction l; simpl; intros; try contradiction. destruct a as [pc1 instr1]; simpl in *. inv H0. monadInv H. inv_incr. destruct (peq pc pc1). - subst. @@ -383,15 +348,13 @@ Proof. + destruct o with pc1; destruct H16; simpl in *; rewrite AssocMap.gss in H14; eauto; congruence. + destruct o0 with pc1; destruct H16; simpl in *; rewrite AssocMap.gss in H14; eauto; congruence. - + inversion H2. inversion H14. unfold nonblock. assert (HST: st_st s4 = st_st s2) by congruence. rewrite HST. - constructor. eapply translate_instr_tr_op. apply EQ1. - eapply in_map with (f := fst) in H14. contradiction. + + inversion H2. inversion H14. unfold nonblock. replace (st_st s4) with (st_st s2) by congruence. + econstructor. apply EQ1. eapply in_map with (f := fst) in H14. contradiction. + destruct o with pc1; destruct H16; simpl in *; rewrite AssocMap.gss in H14; eauto; congruence. + destruct o0 with pc1; destruct H16; simpl in *; rewrite AssocMap.gss in H14; eauto; congruence. - + inversion H2. inversion H14. rewrite <- e2. assert (HST: st_st s2 = st_st s0) by congruence. - rewrite HST. econstructor. apply EQ1. - eapply in_map with (f := fst) in H14. contradiction. + + inversion H2. inversion H14. rewrite <- e2. replace (st_st s2) with (st_st s0) by congruence. + econstructor. apply EQ1. eapply in_map with (f := fst) in H14. contradiction. + destruct o with pc1; destruct H11; simpl in *; rewrite AssocMap.gss in H9; eauto; congruence. + destruct o0 with pc1; destruct H11; simpl in *; rewrite AssocMap.gss in H9; eauto; congruence. @@ -409,6 +372,12 @@ Proof. eauto with htlspec. * apply in_map with (f := fst) in H2. contradiction. + + destruct o with pc1; destruct H16; simpl in *; rewrite AssocMap.gss in H14; eauto; congruence. + + destruct o0 with pc1; destruct H16; simpl in *; rewrite AssocMap.gss in H14; eauto; congruence. + + inversion H2. + * inversion H14. constructor. congruence. + * apply in_map with (f := fst) in H14. contradiction. + + destruct o with pc1; destruct H11; simpl in *; rewrite AssocMap.gss in H9; eauto; congruence. + destruct o0 with pc1; destruct H11; simpl in *; rewrite AssocMap.gss in H9; eauto; congruence. + inversion H2. @@ -428,8 +397,8 @@ Proof. - eapply IHl. apply EQ0. assumption. destruct H2. inversion H2. subst. contradiction. intros. specialize H1 with pc0 instr0. destruct H1. tauto. trivial. - destruct H2. inv H2. contradiction. assumption. assumption.*) -Admitted. + destruct H2. inv H2. contradiction. assumption. assumption. +Qed. Hint Resolve iter_expand_instr_spec : htlspec. Theorem transl_module_correct : -- cgit From a67fb83021f3e5d7ade972ff329ab6c3c4b23620 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Wed, 24 Jun 2020 17:15:22 +0100 Subject: Finish ILoad proof with some assumptions: * EXPR_OK: Yann to work on this. * READ_BOUNDS: To axiomise (or find a better solution). * 32-bit range of register values. --- src/common/IntegerExtra.v | 347 ++++++++++++++++++++---------------------- src/translation/HTLgen.v | 6 +- src/translation/HTLgenproof.v | 36 ++++- 3 files changed, 197 insertions(+), 192 deletions(-) diff --git a/src/common/IntegerExtra.v b/src/common/IntegerExtra.v index ad01015..2f9aae6 100644 --- a/src/common/IntegerExtra.v +++ b/src/common/IntegerExtra.v @@ -11,6 +11,96 @@ Local Open Scope Z_scope. Module PtrofsExtra. + Ltac ptrofs_mod_match m := + match goal with + | [ H : ?x = 0 |- context[?x] ] => rewrite H + | [ _ : _ |- context[_ - 0] ] => rewrite Z.sub_0_r + | [ _ : _ |- context[0 + _] ] => rewrite Z.add_0_l + | [ _ : _ |- context[_ + 0] ] => rewrite Z.add_0_r + | [ _ : _ |- context[0 * _] ] => rewrite Z.mul_0_l + | [ _ : _ |- context[_ * 0] ] => rewrite Z.mul_0_r + | [ _ : _ |- context[?m mod ?m] ] => + rewrite Z_mod_same_full with (a := m) + | [ _ : _ |- context[0 mod _] ] => + rewrite Z.mod_0_l + | [ _ : _ |- context[(_ mod ?m) mod ?m] ] => + rewrite Zmod_mod + | [ _ : _ |- context[(_ mod Ptrofs.modulus) mod m ] ] => + rewrite <- Zmod_div_mod; + try (simplify; lia || assumption) + + | [ _ : _ |- context[Ptrofs.modulus mod m] ] => + rewrite Zdivide_mod with (a := Ptrofs.modulus); + try (lia || assumption) + + | [ _ : _ |- context[Ptrofs.signed ?a mod Ptrofs.modulus] ] => + rewrite Z.mod_small with (a := Ptrofs.signed a) (b := Ptrofs.modulus) + + | [ _ : _ |- context[(?x - ?y) mod ?m] ] => + rewrite Zminus_mod with (a := x) (b := y) (n := m) + + | [ _ : _ |- context[((?x + ?y) mod ?m) + _] ] => + rewrite Zplus_mod with (a := x) (b := y) (n := m) + | [ _ : _ |- context[(?x + ?y) mod ?m] ] => + rewrite Zplus_mod with (a := x) (b := y) (n := m) + + | [ _ : _ |- context[(?x * ?y) mod ?m] ] => + rewrite Zmult_mod with (a := x) (b := y) (n := m) + end. + + Ltac ptrofs_mod_tac m := + repeat (ptrofs_mod_match m); lia. + + Lemma of_int_mod : + forall x m, + Int.signed x mod m = 0 -> + Ptrofs.signed (Ptrofs.of_int x) mod m = 0. + Proof. + intros. + pose proof (Integers.Ptrofs.agree32_of_int eq_refl x) as A. + pose proof Ptrofs.agree32_signed. + apply H0 in A; try reflexivity. + rewrite A. assumption. + Qed. + + Lemma mul_mod : + forall x y m, + 0 < m -> + (m | Ptrofs.modulus) -> + Ptrofs.signed x mod m = 0 -> + Ptrofs.signed y mod m = 0 -> + (Ptrofs.signed (Ptrofs.mul x y)) mod m = 0. + Proof. + intros. unfold Ptrofs.mul. + rewrite Ptrofs.signed_repr_eq. + + repeat match goal with + | [ _ : _ |- context[if ?x then _ else _] ] => destruct x + | [ _ : _ |- context[_ mod Ptrofs.modulus mod m] ] => + rewrite <- Zmod_div_mod; try lia; try assumption + | [ _ : _ |- context[Ptrofs.unsigned _] ] => rewrite Ptrofs.unsigned_signed + end; try (simplify; lia); ptrofs_mod_tac m. + Qed. + + Lemma add_mod : + forall x y m, + 0 < m -> + (m | Ptrofs.modulus) -> + Ptrofs.signed x mod m = 0 -> + Ptrofs.signed y mod m = 0 -> + (Ptrofs.signed (Ptrofs.add x y)) mod m = 0. + Proof. + intros. unfold Ptrofs.add. + rewrite Ptrofs.signed_repr_eq. + + repeat match goal with + | [ _ : _ |- context[if ?x then _ else _] ] => destruct x + | [ _ : _ |- context[_ mod Ptrofs.modulus mod m] ] => + rewrite <- Zmod_div_mod; try lia; try assumption + | [ _ : _ |- context[Ptrofs.unsigned _] ] => rewrite Ptrofs.unsigned_signed + end; try (simplify; lia); ptrofs_mod_tac m. + Qed. + Lemma mul_divs : forall x y, 0 <= Ptrofs.signed y -> @@ -31,24 +121,6 @@ Module PtrofsExtra. congruence. Qed. - Lemma Z_div_distr_add : - forall x y z, - x mod z = 0 -> - y mod z = 0 -> - z <> 0 -> - x / z + y / z = (x + y) / z. - Proof. - intros. - - assert ((x + y) mod z = 0). - { rewrite <- Z.add_mod_idemp_l; try assumption. - rewrite H. assumption. } - - rewrite <- Z.mul_cancel_r with (p := z); try assumption. - rewrite Z.mul_add_distr_r. - repeat rewrite ZLib.div_mul_undo; lia. - Qed. - Lemma mul_unsigned : forall x y, Ptrofs.mul x y = @@ -58,178 +130,85 @@ Module PtrofsExtra. apply Ptrofs.eqm_samerepr. apply Ptrofs.eqm_mult; exists 0; lia. Qed. - - Lemma mul_repr : - forall x y, - Ptrofs.min_signed <= y <= Ptrofs.max_signed -> - Ptrofs.min_signed <= x <= Ptrofs.max_signed -> - Ptrofs.mul (Ptrofs.repr y) (Ptrofs.repr x) = Ptrofs.repr (x * y). - Proof. - intros; unfold Ptrofs.mul. - destruct (Z_ge_lt_dec x 0); destruct (Z_ge_lt_dec y 0). - - - f_equal. - repeat rewrite Ptrofs.unsigned_repr_eq. - repeat rewrite Z.mod_small; simplify; lia. - - - assert (Ptrofs.lt (Ptrofs.repr y) Ptrofs.zero = true). - { - unfold Ptrofs.lt. - rewrite Ptrofs.signed_repr; auto. - rewrite Ptrofs.signed_zero. - destruct (zlt y 0); try lia; auto. - } - - rewrite Ptrofs.unsigned_signed with (n := Ptrofs.repr y). - rewrite H1. - rewrite Ptrofs.signed_repr; auto. - rewrite Ptrofs.unsigned_repr_eq. - rewrite Z.mod_small; simplify; try lia. - rewrite Z.mul_add_distr_r. - apply Ptrofs.eqm_samerepr. - exists x. simplify. lia. - - - assert (Ptrofs.lt (Ptrofs.repr x) Ptrofs.zero = true). - { - unfold Ptrofs.lt. - rewrite Ptrofs.signed_repr; auto. - rewrite Ptrofs.signed_zero. - destruct (zlt x 0); try lia; auto. - } - - rewrite Ptrofs.unsigned_signed with (n := Ptrofs.repr x). - rewrite H1. - rewrite Ptrofs.signed_repr; auto. - rewrite Ptrofs.unsigned_repr_eq. - rewrite Z.mod_small; simplify; try lia. - rewrite Z.mul_add_distr_l. - apply Ptrofs.eqm_samerepr. - exists y. simplify. lia. - - - assert (Ptrofs.lt (Ptrofs.repr x) Ptrofs.zero = true). - { - unfold Ptrofs.lt. - rewrite Ptrofs.signed_repr; auto. - rewrite Ptrofs.signed_zero. - destruct (zlt x 0); try lia; auto. - } - assert (Ptrofs.lt (Ptrofs.repr y) Ptrofs.zero = true). - { - unfold Ptrofs.lt. - rewrite Ptrofs.signed_repr; auto. - rewrite Ptrofs.signed_zero. - destruct (zlt y 0); try lia; auto. - } - rewrite Ptrofs.unsigned_signed with (n := Ptrofs.repr x). - rewrite Ptrofs.unsigned_signed with (n := Ptrofs.repr y). - rewrite H2. - rewrite H1. - repeat rewrite Ptrofs.signed_repr; auto. - replace ((y + Ptrofs.modulus) * (x + Ptrofs.modulus)) with - (x * y + (x + y + Ptrofs.modulus) * Ptrofs.modulus) by lia. - apply Ptrofs.eqm_samerepr. - exists (x + y + Ptrofs.modulus). lia. - Qed. End PtrofsExtra. Module IntExtra. - Lemma mul_unsigned : - forall x y, - Int.mul x y = - Int.repr (Int.unsigned x * Int.unsigned y). + + Ltac int_mod_match m := + match goal with + | [ H : ?x = 0 |- context[?x] ] => rewrite H + | [ _ : _ |- context[_ - 0] ] => rewrite Z.sub_0_r + | [ _ : _ |- context[0 + _] ] => rewrite Z.add_0_l + | [ _ : _ |- context[_ + 0] ] => rewrite Z.add_0_r + | [ _ : _ |- context[0 * _] ] => rewrite Z.mul_0_l + | [ _ : _ |- context[_ * 0] ] => rewrite Z.mul_0_r + | [ _ : _ |- context[?m mod ?m] ] => + rewrite Z_mod_same_full with (a := m) + | [ _ : _ |- context[0 mod _] ] => + rewrite Z.mod_0_l + | [ _ : _ |- context[(_ mod ?m) mod ?m] ] => + rewrite Zmod_mod + | [ _ : _ |- context[(_ mod Int.modulus) mod m ] ] => + rewrite <- Zmod_div_mod; + try (simplify; lia || assumption) + + | [ _ : _ |- context[Int.modulus mod m] ] => + rewrite Zdivide_mod with (a := Int.modulus); + try (lia || assumption) + + | [ _ : _ |- context[Int.signed ?a mod Int.modulus] ] => + rewrite Z.mod_small with (a := Int.signed a) (b := Int.modulus) + + | [ _ : _ |- context[(?x - ?y) mod ?m] ] => + rewrite Zminus_mod with (a := x) (b := y) (n := m) + + | [ _ : _ |- context[((?x + ?y) mod ?m) + _] ] => + rewrite Zplus_mod with (a := x) (b := y) (n := m) + | [ _ : _ |- context[(?x + ?y) mod ?m] ] => + rewrite Zplus_mod with (a := x) (b := y) (n := m) + + | [ _ : _ |- context[(?x * ?y) mod ?m] ] => + rewrite Zmult_mod with (a := x) (b := y) (n := m) + end. + + Ltac int_mod_tac m := + repeat (int_mod_match m); lia. + + Lemma mul_mod : + forall x y m, + 0 < m -> + (m | Int.modulus) -> + Int.signed x mod m = 0 -> + Int.signed y mod m = 0 -> + (Int.signed (Int.mul x y)) mod m = 0. Proof. - intros; unfold Int.mul. - apply Int.eqm_samerepr. - apply Int.eqm_mult; exists 0; lia. + intros. unfold Int.mul. + rewrite Int.signed_repr_eq. + + repeat match goal with + | [ _ : _ |- context[if ?x then _ else _] ] => destruct x + | [ _ : _ |- context[_ mod Int.modulus mod m] ] => + rewrite <- Zmod_div_mod; try lia; try assumption + | [ _ : _ |- context[Int.unsigned _] ] => rewrite Int.unsigned_signed + end; try (simplify; lia); int_mod_tac m. Qed. - Lemma mul_repr : - forall x y, - Int.min_signed <= y <= Int.max_signed -> - Int.min_signed <= x <= Int.max_signed -> - Int.mul (Int.repr y) (Int.repr x) = Int.repr (x * y). + Lemma add_mod : + forall x y m, + 0 < m -> + (m | Int.modulus) -> + Int.signed x mod m = 0 -> + Int.signed y mod m = 0 -> + (Int.signed (Int.add x y)) mod m = 0. Proof. - intros; unfold Int.mul. - destruct (Z_ge_lt_dec x 0); destruct (Z_ge_lt_dec y 0). - - - f_equal. - repeat rewrite Int.unsigned_repr_eq. - repeat rewrite Z.mod_small; simplify; lia. - - - assert (Int.lt (Int.repr y) Int.zero = true). - { - unfold Int.lt. - rewrite Int.signed_repr; auto. - rewrite Int.signed_zero. - destruct (zlt y 0); try lia; auto. - } - - rewrite Int.unsigned_signed with (n := Int.repr y). - rewrite H1. - rewrite Int.signed_repr; auto. - rewrite Int.unsigned_repr_eq. - rewrite Z.mod_small; simplify; try lia. - rewrite Z.mul_add_distr_r. - apply Int.eqm_samerepr. - exists x. simplify. lia. - - - assert (Int.lt (Int.repr x) Int.zero = true). - { - unfold Int.lt. - rewrite Int.signed_repr; auto. - rewrite Int.signed_zero. - destruct (zlt x 0); try lia; auto. - } - - rewrite Int.unsigned_signed with (n := Int.repr x). - rewrite H1. - rewrite Int.signed_repr; auto. - rewrite Int.unsigned_repr_eq. - rewrite Z.mod_small; simplify; try lia. - rewrite Z.mul_add_distr_l. - apply Int.eqm_samerepr. - exists y. simplify. lia. - - - assert (Int.lt (Int.repr x) Int.zero = true). - { - unfold Int.lt. - rewrite Int.signed_repr; auto. - rewrite Int.signed_zero. - destruct (zlt x 0); try lia; auto. - } - assert (Int.lt (Int.repr y) Int.zero = true). - { - unfold Int.lt. - rewrite Int.signed_repr; auto. - rewrite Int.signed_zero. - destruct (zlt y 0); try lia; auto. - } - rewrite Int.unsigned_signed with (n := Int.repr x). - rewrite Int.unsigned_signed with (n := Int.repr y). - rewrite H2. - rewrite H1. - repeat rewrite Int.signed_repr; auto. - replace ((y + Int.modulus) * (x + Int.modulus)) with - (x * y + (x + y + Int.modulus) * Int.modulus) by lia. - apply Int.eqm_samerepr. - exists (x + y + Int.modulus). lia. + intros. unfold Int.add. + rewrite Int.signed_repr_eq. + + repeat match goal with + | [ _ : _ |- context[if ?x then _ else _] ] => destruct x + | [ _ : _ |- context[_ mod Int.modulus mod m] ] => + rewrite <- Zmod_div_mod; try lia; try assumption + | [ _ : _ |- context[Int.unsigned _] ] => rewrite Int.unsigned_signed + end; try (simplify; lia); int_mod_tac m. Qed. End IntExtra. - -Lemma mul_of_int : - forall x y, - 0 <= x < Integers.Ptrofs.modulus -> - Integers.Ptrofs.mul (Integers.Ptrofs.repr x) (Integers.Ptrofs.of_int y) = - Integers.Ptrofs.of_int (Integers.Int.mul (Integers.Int.repr x) y). -Proof. - intros. - pose proof (Integers.Ptrofs.agree32_of_int eq_refl y) as A. - pose proof (Integers.Ptrofs.agree32_to_int eq_refl (Integers.Ptrofs.repr x)) as B. - exploit Integers.Ptrofs.agree32_mul; [> reflexivity | exact B | exact A | intro C]. - unfold Integers.Ptrofs.to_int in C. - unfold Integers.Ptrofs.of_int in C. - rewrite Integers.Ptrofs.unsigned_repr_eq in C. - rewrite Z.mod_small in C; auto. - symmetry. - apply Integers.Ptrofs.agree32_of_int_eq; auto. -Qed. diff --git a/src/translation/HTLgen.v b/src/translation/HTLgen.v index 92e40f5..357d487 100644 --- a/src/translation/HTLgen.v +++ b/src/translation/HTLgen.v @@ -260,6 +260,10 @@ Definition translate_eff_addressing (a: Op.addressing) (args: list reg) : mon ex if (check_address_parameter scale) && (check_address_parameter offset) then ret (Vbinop Vadd (boplitz Vmul r1 scale) (Vlit (ZToValue 32 offset))) else error (Errors.msg "Veriloggen: translate_eff_addressing address misaligned") + | Op.Aindexed2 offset, r1::r2::nil => + if (check_address_parameter offset) + then ret (Vbinop Vadd (Vvar r1) (boplitz Vadd r2 offset)) + else error (Errors.msg "Veriloggen: translate_eff_addressing address misaligned") | Op.Aindexed2scaled scale offset, r1::r2::nil => (* Typical for dynamic array addressing *) if (check_address_parameter scale) && (check_address_parameter offset) then ret (Vbinop Vadd (boplitz Vadd r1 offset) (boplitz Vmul r2 scale)) @@ -363,7 +367,7 @@ Definition translate_arr_access (mem : AST.memory_chunk) (addr : Op.addressing) (ZToValue 32 4))) else error (Errors.msg "Veriloggen: translate_arr_access address misaligned") | Mint32, Op.Ainstack a, nil => (* We need to be sure that the base address is aligned *) - let a := Integers.Ptrofs.unsigned a in + let a := Integers.Ptrofs.signed a in if (check_address_parameter a) then ret (Vvari stack (Vlit (ZToValue 32 (a / 4)))) else error (Errors.msg "Veriloggen: eff_addressing misaligned stack offset") diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index 8e97c58..a502453 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -525,8 +525,16 @@ Section CORRECTNESS. assert (0 <= Integers.Ptrofs.signed OFFSET) as READ_BOUND_LOW by admit. assert (Integers.Ptrofs.signed OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH by admit. - (** Modular Preservation proof *) - assert (Integers.Ptrofs.signed OFFSET mod 4 = 0) as MOD_PRESERVE by admit. + (** Modular preservation proof *) + assert (Integers.Ptrofs.signed OFFSET mod 4 = 0) as MOD_PRESERVE. + { rewrite HeqOFFSET. + apply PtrofsExtra.add_mod; simplify; try lia. + exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *) + rewrite Integers.Ptrofs.signed_repr; try assumption. + admit. (* FIXME: Register bounds. *) + apply PtrofsExtra.of_int_mod. + rewrite Integers.Int.signed_repr; simplify; try split; try assumption. + } (** Normalisation proof *) assert (Integers.Ptrofs.repr @@ -734,8 +742,22 @@ Section CORRECTNESS. assert (0 <= Integers.Ptrofs.signed OFFSET) as READ_BOUND_LOW by admit. assert (Integers.Ptrofs.signed OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH by admit. - (** Modular Preservation proof *) - assert (Integers.Ptrofs.signed OFFSET mod 4 = 0) as MOD_PRESERVE by admit. + (** Modular preservation proof *) + assert (Integers.Ptrofs.signed OFFSET mod 4 = 0) as MOD_PRESERVE. + { rewrite HeqOFFSET. + apply PtrofsExtra.add_mod; simplify; try lia. + exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *) + rewrite Integers.Ptrofs.signed_repr; try assumption. + admit. (* FIXME: Register bounds. *) + apply PtrofsExtra.of_int_mod. + apply IntExtra.add_mod; simplify; try lia. + exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *) + apply IntExtra.mul_mod; simplify; try lia. + exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *) + admit. (* FIXME: Register bounds. *) + rewrite Integers.Int.signed_repr; simplify; try split; try assumption. + rewrite Integers.Int.signed_repr; simplify; try split; try assumption. + } (** Normalisation proof *) assert (Integers.Ptrofs.repr @@ -918,8 +940,8 @@ Section CORRECTNESS. assert (0 <= Integers.Ptrofs.signed OFFSET) as READ_BOUND_LOW by admit. assert (Integers.Ptrofs.signed OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH by admit. - (** Modular Preservation proof *) - assert (Integers.Ptrofs.signed OFFSET mod 4 = 0) as MOD_PRESERVE by admit. + (** Modular preservation proof *) + rename H8 into MOD_PRESERVE. (** Normalisation proof *) assert (Integers.Ptrofs.repr @@ -1006,7 +1028,7 @@ Section CORRECTNESS. OFFSET (Integers.Ptrofs.repr 4))) = - valueToNat (ZToValue 32 (Integers.Ptrofs.unsigned OFFSET / 4))) + valueToNat (ZToValue 32 (Integers.Ptrofs.signed OFFSET / 4))) as EXPR_OK by admit. rewrite <- EXPR_OK. rewrite NORMALISE in I. -- cgit From 9e29a351fd0928130fa0b67dc47b67cdf989e4b7 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Wed, 24 Jun 2020 17:53:36 +0100 Subject: Fixes to make develop compile --- src/translation/HTLgen.v | 2 +- src/translation/HTLgenproof.v | 2 +- src/translation/HTLgenspec.v | 2 +- src/verilog/Value.v | 8 +++----- 4 files changed, 6 insertions(+), 8 deletions(-) diff --git a/src/translation/HTLgen.v b/src/translation/HTLgen.v index 66170bc..eb2ddda 100644 --- a/src/translation/HTLgen.v +++ b/src/translation/HTLgen.v @@ -559,7 +559,7 @@ Definition transl_program (p: RTL.program) : Errors.res HTL.program := (fun i v => Errors.OK v) p. *) -Definition main_is_internal {A B : Type} (p : AST.program A B) : bool := +Definition main_is_internal (p : RTL.program) : bool := let ge := Globalenvs.Genv.globalenv p in match Globalenvs.Genv.find_symbol ge p.(AST.prog_main) with | Some b => diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index 9b59269..724ac0a 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -393,7 +393,7 @@ Section CORRECTNESS. 2: { reflexivity. } rewrite AssocMap.gss. unfold Verilog.merge_arr. - setoid_rewrite H1. + setoid_rewrite H3. reflexivity. rewrite combine_length. diff --git a/src/translation/HTLgenspec.v b/src/translation/HTLgenspec.v index 7d89a76..c617174 100644 --- a/src/translation/HTLgenspec.v +++ b/src/translation/HTLgenspec.v @@ -140,7 +140,7 @@ Inductive tr_instr (fin rtrn st stk : reg) : RTL.instruction -> stmnt -> stmnt - forall mem addr args s s' i c src n, translate_arr_access mem addr args stk s = OK c s' i -> tr_instr fin rtrn st stk (RTL.Istore mem addr args src n) (Vnonblock c (Vvar src)) - (state_goto st n). + (state_goto st n) | tr_instr_Ijumptable : forall cexpr tbl r, cexpr = tbl_to_case_expr st tbl -> diff --git a/src/verilog/Value.v b/src/verilog/Value.v index bde98b8..52a87e3 100644 --- a/src/verilog/Value.v +++ b/src/verilog/Value.v @@ -372,7 +372,7 @@ Lemma boolToValue_ValueToBool : valueToBool (boolToValue 32 b) = b. Proof. destruct b; unfold valueToBool, boolToValue; simpl; trivial. Qed. -Lemma ZToValue_valueToNat : +(*Lemma ZToValue_valueToNat : forall x sz, sz > 0 -> (x < 2^(Z.of_nat sz))%Z -> @@ -384,7 +384,5 @@ Proof. unfold Z.of_nat in *. destruct sz eqn:?. omega. simpl in H0. rewrite <- Pos2Z.inj_pow_pos in H0. Search (Z.pos _ < Z.pos _)%Z. Search Pos.to_nat (_ < _). (* Pos2Nat.inj_lt *) - Search Pos.to_nat. - (* Pos2Nat.is_succ: forall p : positive, exists n : nat, Pos.to_nat p = S n *) - Search S Pos.to_nat. Search (_ < _ <-> _)%positive. assert (exists p, Pos.to_nat p = S n). - econstructor. Search Pos2Nat.is_succ. Search Pos.of_succ_nat. + Search "inj" positive nat. +*) -- cgit From bb80bc5d196665498f7b365e9e26468ed5999ea9 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Wed, 24 Jun 2020 20:09:13 +0100 Subject: HTLgenproof passing. --- src/translation/HTLgen.v | 4 - src/translation/HTLgenproof.v | 173 +++++++++++++++++++++++++++++++----------- 2 files changed, 127 insertions(+), 50 deletions(-) diff --git a/src/translation/HTLgen.v b/src/translation/HTLgen.v index e637d6f..54ad77a 100644 --- a/src/translation/HTLgen.v +++ b/src/translation/HTLgen.v @@ -395,10 +395,6 @@ Definition translate_arr_access (mem : AST.memory_chunk) (addr : Op.addressing) if (check_address_parameter off) then ret (Vvari stack (Vbinop Vdiv (boplitz Vadd r1 off) (Vlit (ZToValue 32 4)))) else error (Errors.msg "HTLgen: translate_arr_access address misaligned") - | Mint32, Op.Ascaled scale offset, r1::nil => - if (check_address_parameter scale) && (check_address_parameter offset) - then ret (Vvari stack (Vbinop Vdiv (Vbinop Vadd (boplitz Vmul r1 scale) (Vlit (ZToValue 32 offset))) (Vlit (ZToValue 32 4)))) - else error (Errors.msg "HTLgen: translate_arr_access address misaligned") | Mint32, Op.Aindexed2scaled scale offset, r1::r2::nil => (* Typical for dynamic array addressing *) if (check_address_parameter scale) && (check_address_parameter offset) then ret (Vvari stack diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index 24191ae..cee04e9 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -57,18 +57,6 @@ Inductive match_arrs (m : HTL.module) (f : RTL.function) (sp : Values.val) (mem (Option.join (array_get_error (Z.to_nat ptr) stack)))) -> match_arrs m f sp mem asa. -Inductive match_stacks (mem : mem) : list RTL.stackframe -> list HTL.stackframe -> Prop := -| match_stacks_nil : - match_stacks mem nil nil -| match_stacks_cons : - forall cs lr r f sp pc rs m asr asa - (TF : tr_module f m) - (ST: match_stacks mem cs lr) - (MA: match_assocmaps f rs asr) - (MARR : match_arrs m f sp mem asa), - match_stacks mem (RTL.Stackframe r f sp pc rs :: cs) - (HTL.Stackframe r m pc asr asa :: lr). - Definition stack_based (v : Values.val) (sp : Values.block) : Prop := match v with | Values.Vptr sp' off => sp' = sp @@ -88,6 +76,21 @@ Definition arr_stack_based_pointers (spb : Values.block) (m : mem) (stack_length (Values.Val.offset_ptr sp (Integers.Ptrofs.repr (4 * ptr))))) spb. +Inductive match_stacks (mem : mem) : list RTL.stackframe -> list HTL.stackframe -> Prop := +| match_stacks_nil : + match_stacks mem nil nil +| match_stacks_cons : + forall cs lr r f sp sp' pc rs m asr asa + (TF : tr_module f m) + (ST: match_stacks mem cs lr) + (MA: match_assocmaps f rs asr) + (MARR : match_arrs m f sp mem asa) + (SP : sp = Values.Vptr sp' (Integers.Ptrofs.repr 0)) + (RSBP: reg_stack_based_pointers sp' rs) + (ASBP: arr_stack_based_pointers sp' mem (f.(RTL.fn_stacksize)) sp), + match_stacks mem (RTL.Stackframe r f sp pc rs :: cs) + (HTL.Stackframe r m pc asr asa :: lr). + Inductive match_states : RTL.state -> HTL.state -> Prop := | match_state : forall asa asr sf f sp sp' rs mem m st res (MASSOC : match_assocmaps f rs asr) @@ -707,11 +710,6 @@ Section CORRECTNESS. rewrite Registers.Regmap.gso; auto. - + (* FIXME: Why is this degenerate? Should we support this mode? *) - unfold Op.eval_addressing in H0. - destruct (Archi.ptr64) eqn:ARCHI; simplify. - destruct (Registers.Regmap.get r0 rs) eqn:EQr0; discriminate. - + (** Preamble *) invert MARR. simplify. @@ -1115,12 +1113,6 @@ Section CORRECTNESS. + admit. + admit. - (* + eexists. split. *) - (* eapply Smallstep.plus_one. *) - (* eapply HTL.step_module; eauto. *) - (* econstructor. econstructor. econstructor. simplify. *) - + admit. - - eexists. split. apply Smallstep.plus_one. eapply HTL.step_module; eauto. eapply Verilog.stmnt_runp_Vnonblock_reg with @@ -1139,37 +1131,95 @@ Section CORRECTNESS. constructor. apply boolToValue_ValueToBool. constructor. + unfold Verilog.merge_regs. unfold_merge. apply AssocMap.gss. destruct b. rewrite assumption_32bit. - apply match_state. + simplify. + apply match_state with (sp' := sp'); eauto. + unfold Verilog.merge_regs. unfold_merge. apply regs_lessdef_add_greater. apply greater_than_max_func. - assumption. assumption. + assumption. - unfold state_st_wf. intros. inversion H1. - subst. unfold_merge. + unfold state_st_wf. intros. + invert H3. + unfold Verilog.merge_regs. unfold_merge. apply AssocMap.gss. - assumption. + (** Match arrays *) + invert MARR. simplify. + econstructor. + repeat split; simplify. + unfold HTL.empty_stack. + simplify. + unfold Verilog.merge_arrs. + + rewrite AssocMap.gcombine. + 2: { reflexivity. } + rewrite AssocMap.gss. + unfold Verilog.merge_arr. + setoid_rewrite H4. + reflexivity. + + rewrite combine_length. + unfold arr_repeat. simplify. + rewrite list_repeat_len. + reflexivity. + + unfold arr_repeat. simplify. + rewrite list_repeat_len. + congruence. + + intros. + erewrite array_get_error_equal. + eauto. apply combine_none. assumption. rewrite assumption_32bit. - apply match_state. - unfold_merge. - apply regs_lessdef_add_greater. apply greater_than_max_func. assumption. + apply match_state with (sp' := sp'); eauto. + unfold Verilog.merge_regs. unfold_merge. + apply regs_lessdef_add_greater. apply greater_than_max_func. assumption. - unfold state_st_wf. intros. inversion H1. - subst. unfold_merge. + unfold state_st_wf. intros. + invert H1. + unfold Verilog.merge_regs. unfold_merge. apply AssocMap.gss. - assumption. + (** Match arrays *) + invert MARR. simplify. + econstructor. + repeat split; simplify. + unfold HTL.empty_stack. + simplify. + unfold Verilog.merge_arrs. + + rewrite AssocMap.gcombine. + 2: { reflexivity. } + rewrite AssocMap.gss. + unfold Verilog.merge_arr. + setoid_rewrite H2. + reflexivity. + + rewrite combine_length. + unfold arr_repeat. simplify. + rewrite list_repeat_len. + reflexivity. + + unfold arr_repeat. simplify. + rewrite list_repeat_len. + congruence. + + intros. + erewrite array_get_error_equal. + eauto. apply combine_none. assumption. - admit. + - (* Return *) econstructor. split. eapply Smallstep.plus_two. @@ -1184,6 +1234,7 @@ Section CORRECTNESS. constructor. constructor. + unfold Verilog.merge_regs. unfold_merge. simpl. rewrite AssocMap.gso. rewrite AssocMap.gso. @@ -1191,6 +1242,7 @@ Section CORRECTNESS. apply st_greater_than_res. apply st_greater_than_res. apply HTL.step_finish. + unfold Verilog.merge_regs. unfold_merge; simpl. rewrite AssocMap.gso. apply AssocMap.gss. @@ -1214,6 +1266,7 @@ Section CORRECTNESS. constructor. econstructor; simpl; trivial. apply Verilog.erun_Vvar. trivial. + unfold Verilog.merge_regs. unfold_merge. simpl. rewrite AssocMap.gso. rewrite AssocMap.gso. @@ -1221,6 +1274,7 @@ Section CORRECTNESS. apply st_greater_than_res. apply st_greater_than_res. trivial. apply HTL.step_finish. + unfold Verilog.merge_regs. unfold_merge. rewrite AssocMap.gso. apply AssocMap.gss. @@ -1238,21 +1292,26 @@ Section CORRECTNESS. - inversion MSTATE; subst. inversion TF; subst. econstructor. split. apply Smallstep.plus_one. - eapply HTL.step_call. simpl. + eapply HTL.step_call. simplify. - constructor. apply regs_lessdef_add_greater. + apply match_state with (sp' := stk); eauto. + + apply regs_lessdef_add_greater. apply greater_than_max_func. - apply init_reg_assoc_empty. assumption. unfold state_st_wf. - intros. inv H1. apply AssocMap.gss. constructor. + apply init_reg_assoc_empty. + unfold state_st_wf. + intros. inv H3. apply AssocMap.gss. constructor. - econstructor; simpl. - reflexivity. - rewrite AssocMap.gss. reflexivity. + econstructor. simplify. + repeat split. unfold HTL.empty_stack. + simplify. apply AssocMap.gss. + unfold arr_repeat. simplify. + apply list_repeat_len. intros. destruct (Mem.load AST.Mint32 m' stk (Integers.Ptrofs.unsigned (Integers.Ptrofs.add Integers.Ptrofs.zero - (Integers.Ptrofs.repr ptr)))) eqn:LOAD. + (Integers.Ptrofs.repr (4 * ptr))))) eqn:LOAD. pose proof Mem.load_alloc_same as LOAD_ALLOC. pose proof H as ALLOC. eapply LOAD_ALLOC in ALLOC. @@ -1261,21 +1320,43 @@ Section CORRECTNESS. repeat constructor. constructor. - - inversion MSTATE; subst. inversion MS; subst. econstructor. + unfold reg_stack_based_pointers. intros. + unfold RTL.init_regs; simplify. + destruct (RTL.fn_params f); + rewrite Registers.Regmap.gi; constructor. + + unfold arr_stack_based_pointers. intros. + simplify. + destruct (Mem.load AST.Mint32 m' stk + (Integers.Ptrofs.unsigned (Integers.Ptrofs.add + Integers.Ptrofs.zero + (Integers.Ptrofs.repr (4 * ptr))))) eqn:LOAD. + pose proof Mem.load_alloc_same as LOAD_ALLOC. + pose proof H as ALLOC. + eapply LOAD_ALLOC in ALLOC. + 2: { exact LOAD. } + rewrite ALLOC. + repeat constructor. + constructor. + + - invert MSTATE. invert MS. + econstructor. split. apply Smallstep.plus_one. constructor. - constructor; auto. constructor; auto. apply regs_lessdef_add_match; auto. + constructor; auto. + econstructor; auto. + apply regs_lessdef_add_match; auto. apply regs_lessdef_add_greater. apply greater_than_max_func. auto. unfold state_st_wf. intros. inv H. rewrite AssocMap.gso. rewrite AssocMap.gss. trivial. apply st_greater_than_res. + admit. + Unshelve. exact (AssocMap.empty value). exact (AssocMap.empty value). - exact (AssocMap.empty value). - exact (AssocMap.empty value). (* exact (ZToValue 32 0). *) (* exact (AssocMap.empty value). *) (* exact (AssocMap.empty value). *) -- cgit From 14cd3b8b12e3db17c3842ae9dfdb30ca86b6394c Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Wed, 24 Jun 2020 20:57:54 +0100 Subject: Finish Internal main proof --- src/translation/HTLgenproof.v | 43 +++++++++++++++++++++++++++++-------------- 1 file changed, 29 insertions(+), 14 deletions(-) diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index cee04e9..f07403d 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -1365,16 +1365,29 @@ Section CORRECTNESS. Admitted. Hint Resolve transl_step_correct : htlproof. - Lemma main_tprog_internal : - forall b f, - Globalenvs.Genv.find_symbol tge tprog.(AST.prog_main) = Some b -> - Genv.find_funct_ptr (Genv.globalenv tprog) b = Some (AST.Internal f). - Admitted. Lemma option_inv : forall A x y, @Some A x = Some y -> x = y. - Proof. intros. inversion H. trivial. Qed. + Proof. intros. inversion H. trivial. Qed. + + Lemma main_tprog_internal : + forall b, + Globalenvs.Genv.find_symbol tge tprog.(AST.prog_main) = Some b -> + exists f, Genv.find_funct_ptr (Genv.globalenv tprog) b = Some (AST.Internal f). + Proof. + intros. + destruct TRANSL. unfold main_is_internal in H1. + repeat (unfold_match H1). replace b with b0. + exploit function_ptr_translated; eauto. intros [tf [A B]]. + unfold transl_fundef, AST.transf_partial_fundef, Errors.bind in B. + unfold_match B. inv B. econstructor. apply A. + + apply option_inv. rewrite <- Heqo. rewrite <- H. + rewrite symbols_preserved. replace (AST.prog_main tprog) with (AST.prog_main prog). + trivial. symmetry; eapply Linking.match_program_main; eauto. + Qed. + (* Had to admit proof because currently there is no way to force main to be Internal. *) Lemma transl_initial_states : @@ -1395,21 +1408,23 @@ Section CORRECTNESS. unfold transl_fundef, Errors.bind in B. unfold AST.transf_partial_fundef, Errors.bind in B. repeat (unfold_match B). inversion B. subst. + exploit main_tprog_internal; eauto; intros. + rewrite symbols_preserved. replace (AST.prog_main tprog) with (AST.prog_main prog). + apply Heqo. symmetry; eapply Linking.match_program_main; eauto. + inversion H5. econstructor; split. econstructor. apply (Genv.init_mem_transf_partial TRANSL'); eauto. replace (AST.prog_main tprog) with (AST.prog_main prog). rewrite symbols_preserved; eauto. symmetry; eapply Linking.match_program_main; eauto. - apply main_tprog_internal. replace ge0 with ge in * by auto. - replace b0 with b in *. rewrite symbols_preserved. - replace (AST.prog_main tprog) with (AST.prog_main prog). - assumption. - symmetry; eapply Linking.match_program_main; eauto. - apply option_inv. rewrite <- H0. rewrite <- Heqo. - trivial. + apply H6. constructor. - apply transl_module_correct. eassumption. + apply transl_module_correct. + assert (Some (AST.Internal x) = Some (AST.Internal m)). + replace (AST.fundef HTL.module) with (HTL.fundef). + rewrite <- H6. setoid_rewrite <- A. trivial. + trivial. inversion H7. Qed. Hint Resolve transl_initial_states : htlproof. -- cgit From 445aabbcf63e29d68dd0c98dde7f259af0381591 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 25 Jun 2020 09:09:23 +0100 Subject: Work on Veriloggen proof --- src/translation/HTLgenproof.v | 2 +- src/translation/Veriloggenproof.v | 52 +++++++++++++++++++++++++++++++++++++-- src/verilog/Verilog.v | 7 +++--- 3 files changed, 55 insertions(+), 6 deletions(-) diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index f07403d..77bd4ef 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -1424,7 +1424,7 @@ Section CORRECTNESS. assert (Some (AST.Internal x) = Some (AST.Internal m)). replace (AST.fundef HTL.module) with (HTL.fundef). rewrite <- H6. setoid_rewrite <- A. trivial. - trivial. inversion H7. + trivial. inv H7. assumption. Qed. Hint Resolve transl_initial_states : htlproof. diff --git a/src/translation/Veriloggenproof.v b/src/translation/Veriloggenproof.v index 6c58c56..825cb7e 100644 --- a/src/translation/Veriloggenproof.v +++ b/src/translation/Veriloggenproof.v @@ -16,16 +16,64 @@ * along with this program. If not, see . *) -From compcert Require Import Smallstep. +From compcert Require Import Smallstep Linking. +From coqup Require Import Veriloggen. From coqup Require HTL Verilog. +Definition match_prog (prog : HTL.program) (tprog : Verilog.program) := + match_program (fun cu f tf => tf = transl_fundef f) eq prog tprog. + +Lemma transf_program_match: + forall prog, match_prog prog (transl_program prog). +Proof. + intros. eapply match_transform_program_contextual. auto. +Qed. + +Inductive match_stacks : list HTL.stackframe -> list Verilog.stackframe -> Prop := +| match_stack : + forall res m pc reg_assoc arr_assoc hstk vstk, + match_stacks hstk vstk -> + match_stacks (HTL.Stackframe res m pc reg_assoc arr_assoc :: hstk) + (Verilog.Stackframe res (transl_module m) pc + reg_assoc arr_assoc :: vstk) +| match_stack_nil : match_stacks nil nil. + +Inductive match_states : HTL.state -> Verilog.state -> Prop := +| match_state : + forall m st reg_assoc arr_assoc hstk vstk, + match_stacks hstk vstk -> + match_states (HTL.State hstk m st reg_assoc arr_assoc) + (Verilog.State vstk (transl_module m) st reg_assoc arr_assoc) +| match_returnstate : + forall v hstk vstk, + match_stacks hstk vstk -> + match_states (HTL.Returnstate hstk v) (Verilog.Returnstate vstk v) +| match_initial_call : + forall m, + match_states (HTL.Callstate nil m nil) (Verilog.Callstate nil (transl_module m) nil). + Section CORRECTNESS. Variable prog: HTL.program. Variable tprog: Verilog.program. + Hypothesis TRANSL : match_prog prog tprog. + + Let ge : HTL.genv := Globalenvs.Genv.globalenv prog. + Let tge : Verilog.genv := Globalenvs.Genv.globalenv tprog. + + Theorem transl_step_correct: + forall (S1 : HTL.state) t S2, + HTL.step ge S1 t S2 -> + forall (R1 : Verilog.state), + match_states S1 R1 -> + exists R2, Smallstep.plus Verilog.step tge R1 t R2 /\ match_states S2 R2. + Proof. + induction 1; intros R1 MSTATE. + Theorem transf_program_correct: forward_simulation (HTL.semantics prog) (Verilog.semantics tprog). - Admitted. + End CORRECTNESS. + diff --git a/src/verilog/Verilog.v b/src/verilog/Verilog.v index 7d5e3c0..9c05fc9 100644 --- a/src/verilog/Verilog.v +++ b/src/verilog/Verilog.v @@ -279,7 +279,8 @@ Inductive stackframe : Type := forall (res : reg) (m : module) (pc : node) - (assoc : assocmap), + (reg_assoc : assocmap_reg) + (arr_assoc : assocmap_arr), stackframe. Inductive state : Type := @@ -735,9 +736,9 @@ Inductive step : genv -> state -> Events.trace -> state -> Prop := (init_params args m.(mod_args))) (empty_stack m)) | step_return : - forall g m asr i r sf pc mst, + forall g m asr i r sf pc mst asa, mst = mod_st m -> - step g (Returnstate (Stackframe r m pc asr :: sf) i) Events.E0 + step g (Returnstate (Stackframe r m pc asr asa :: sf) i) Events.E0 (State sf m pc ((asr # mst <- (posToValue 32 pc)) # r <- i) (empty_stack m)). Hint Constructors step : verilog. -- cgit From cf9949a5151aa9ed86554fb31c2a56fad0614a10 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 25 Jun 2020 18:04:49 +0100 Subject: Progress on proof of Veriloggen --- src/translation/Veriloggenproof.v | 33 +++++++++++++++++++-------------- src/verilog/HTL.v | 4 +++- src/verilog/Verilog.v | 2 ++ 3 files changed, 24 insertions(+), 15 deletions(-) diff --git a/src/translation/Veriloggenproof.v b/src/translation/Veriloggenproof.v index 825cb7e..e556c69 100644 --- a/src/translation/Veriloggenproof.v +++ b/src/translation/Veriloggenproof.v @@ -17,8 +17,8 @@ *) From compcert Require Import Smallstep Linking. -From coqup Require Import Veriloggen. -From coqup Require HTL Verilog. +From coqup Require HTL. +From coqup Require Import Coquplib Veriloggen Verilog. Definition match_prog (prog : HTL.program) (tprog : Verilog.program) := match_program (fun cu f tf => tf = transl_fundef f) eq prog tprog. @@ -29,47 +29,52 @@ Proof. intros. eapply match_transform_program_contextual. auto. Qed. -Inductive match_stacks : list HTL.stackframe -> list Verilog.stackframe -> Prop := +Inductive match_stacks : list HTL.stackframe -> list stackframe -> Prop := | match_stack : forall res m pc reg_assoc arr_assoc hstk vstk, match_stacks hstk vstk -> match_stacks (HTL.Stackframe res m pc reg_assoc arr_assoc :: hstk) - (Verilog.Stackframe res (transl_module m) pc + (Stackframe res (transl_module m) pc reg_assoc arr_assoc :: vstk) | match_stack_nil : match_stacks nil nil. -Inductive match_states : HTL.state -> Verilog.state -> Prop := +Inductive match_states : HTL.state -> state -> Prop := | match_state : forall m st reg_assoc arr_assoc hstk vstk, match_stacks hstk vstk -> match_states (HTL.State hstk m st reg_assoc arr_assoc) - (Verilog.State vstk (transl_module m) st reg_assoc arr_assoc) + (State vstk (transl_module m) st reg_assoc arr_assoc) | match_returnstate : forall v hstk vstk, match_stacks hstk vstk -> - match_states (HTL.Returnstate hstk v) (Verilog.Returnstate vstk v) + match_states (HTL.Returnstate hstk v) (Returnstate vstk v) | match_initial_call : forall m, - match_states (HTL.Callstate nil m nil) (Verilog.Callstate nil (transl_module m) nil). + match_states (HTL.Callstate nil m nil) (Callstate nil (transl_module m) nil). Section CORRECTNESS. Variable prog: HTL.program. - Variable tprog: Verilog.program. + Variable tprog: program. Hypothesis TRANSL : match_prog prog tprog. Let ge : HTL.genv := Globalenvs.Genv.globalenv prog. - Let tge : Verilog.genv := Globalenvs.Genv.globalenv tprog. + Let tge : genv := Globalenvs.Genv.globalenv tprog. - Theorem transl_step_correct: + Lemma stmnt_in_case : + exists e st, + + Theorem transl_step_correct : forall (S1 : HTL.state) t S2, HTL.step ge S1 t S2 -> - forall (R1 : Verilog.state), + forall (R1 : state), match_states S1 R1 -> - exists R2, Smallstep.plus Verilog.step tge R1 t R2 /\ match_states S2 R2. + exists R2, Smallstep.plus step tge R1 t R2 /\ match_states S2 R2. Proof. - induction 1; intros R1 MSTATE. + induction 1; intros R1 MSTATE; inv MSTATE; econstructor; split. + - apply Smallstep.plus_one. econstructor. econstructor. + * econstructor. Theorem transf_program_correct: forward_simulation (HTL.semantics prog) (Verilog.semantics tprog). diff --git a/src/verilog/HTL.v b/src/verilog/HTL.v index 0bf5072..a3623f0 100644 --- a/src/verilog/HTL.v +++ b/src/verilog/HTL.v @@ -99,12 +99,14 @@ Inductive state : Type := Inductive step : genv -> state -> Events.trace -> state -> Prop := | step_module : - forall g m st sf ctrl data + forall g m st sf ctrl data ist asr asa basr1 basa1 nasr1 nasa1 basr2 basa2 nasr2 nasa2 asr' asa' f stval pstval, + asr!(m.(mod_st)) = Some ist -> + valueToPos ist = st -> m.(mod_controllogic)!st = Some ctrl -> m.(mod_datapath)!st = Some data -> Verilog.stmnt_runp f diff --git a/src/verilog/Verilog.v b/src/verilog/Verilog.v index 9c05fc9..d476710 100644 --- a/src/verilog/Verilog.v +++ b/src/verilog/Verilog.v @@ -713,6 +713,8 @@ Definition empty_stack (m : module) : assocmap_arr := Inductive step : genv -> state -> Events.trace -> state -> Prop := | step_module : forall asr asa asr' asa' basr1 nasr1 basa1 nasa1 f stval pstval m sf st g, + asr!(m.(mod_st)) = Some ist -> + valueToPos ist = st -> mis_stepp f (mkassociations asr empty_assocmap) (mkassociations asa (empty_stack m)) m.(mod_body) -- cgit From 50ec2fb12454c2bc1f902c955f0b81df71b58c39 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Fri, 26 Jun 2020 09:40:16 +0100 Subject: Fix Verilog semantics and fix order of always blocks --- src/translation/Veriloggen.v | 4 ++-- src/translation/Veriloggenproof.v | 7 ++----- src/verilog/Verilog.v | 2 +- 3 files changed, 5 insertions(+), 8 deletions(-) diff --git a/src/translation/Veriloggen.v b/src/translation/Veriloggen.v index 2b9974b..b550ff9 100644 --- a/src/translation/Veriloggen.v +++ b/src/translation/Veriloggen.v @@ -43,10 +43,10 @@ Definition transl_module (m : HTL.module) : Verilog.module := let case_el_ctrl := transl_list (PTree.elements m.(mod_controllogic)) in let case_el_data := transl_list (PTree.elements m.(mod_datapath)) in let body := - Valways (Vposedge m.(mod_clk)) (Vcase (Vvar m.(mod_st)) case_el_data (Some Vskip)) - :: Valways (Vposedge m.(mod_clk)) (Vcond (Vbinop Veq (Vvar m.(mod_reset)) (ZToValue 1 1)) + Valways (Vposedge m.(mod_clk)) (Vcond (Vbinop Veq (Vvar m.(mod_reset)) (ZToValue 1 1)) (Vnonblock (Vvar m.(mod_st)) (posToValue 32 m.(mod_entrypoint))) (Vcase (Vvar m.(mod_st)) case_el_ctrl (Some Vskip))) + :: Valways (Vposedge m.(mod_clk)) (Vcase (Vvar m.(mod_st)) case_el_data (Some Vskip)) :: (arr_to_Vdeclarr (AssocMap.elements m.(mod_arrdecls)) ++ scl_to_Vdecl (AssocMap.elements m.(mod_scldecls))) in Verilog.mkmodule m.(mod_start) diff --git a/src/translation/Veriloggenproof.v b/src/translation/Veriloggenproof.v index e556c69..db96949 100644 --- a/src/translation/Veriloggenproof.v +++ b/src/translation/Veriloggenproof.v @@ -62,9 +62,6 @@ Section CORRECTNESS. Let ge : HTL.genv := Globalenvs.Genv.globalenv prog. Let tge : genv := Globalenvs.Genv.globalenv tprog. - Lemma stmnt_in_case : - exists e st, - Theorem transl_step_correct : forall (S1 : HTL.state) t S2, HTL.step ge S1 t S2 -> @@ -73,8 +70,8 @@ Section CORRECTNESS. exists R2, Smallstep.plus step tge R1 t R2 /\ match_states S2 R2. Proof. induction 1; intros R1 MSTATE; inv MSTATE; econstructor; split. - - apply Smallstep.plus_one. econstructor. econstructor. - * econstructor. + - apply Smallstep.plus_one. econstructor. eassumption. trivial. + * econstructor. econstructor. Theorem transf_program_correct: forward_simulation (HTL.semantics prog) (Verilog.semantics tprog). diff --git a/src/verilog/Verilog.v b/src/verilog/Verilog.v index d476710..555ddbd 100644 --- a/src/verilog/Verilog.v +++ b/src/verilog/Verilog.v @@ -712,7 +712,7 @@ Definition empty_stack (m : module) : assocmap_arr := Inductive step : genv -> state -> Events.trace -> state -> Prop := | step_module : - forall asr asa asr' asa' basr1 nasr1 basa1 nasa1 f stval pstval m sf st g, + forall asr asa asr' asa' basr1 nasr1 basa1 nasa1 f stval pstval m sf st g ist, asr!(m.(mod_st)) = Some ist -> valueToPos ist = st -> mis_stepp f (mkassociations asr empty_assocmap) -- cgit From 8fda19cb580bda72f374bc2176d7e2efa5cd613b Mon Sep 17 00:00:00 2001 From: James Pollard Date: Sun, 28 Jun 2020 15:06:17 +0100 Subject: Work on proof. --- src/common/Coquplib.v | 3 + src/common/IntegerExtra.v | 50 +++++++ src/translation/HTLgenproof.v | 301 +++++++++++++++++++++++++++++++++++++++++- src/verilog/Array.v | 53 ++++++++ 4 files changed, 405 insertions(+), 2 deletions(-) diff --git a/src/common/Coquplib.v b/src/common/Coquplib.v index b4ca906..b8a02d2 100644 --- a/src/common/Coquplib.v +++ b/src/common/Coquplib.v @@ -117,6 +117,9 @@ Ltac simplify := unfold_constants; simpl in *; Global Opaque Nat.div. Global Opaque Z.mul. +Infix "==nat" := eq_nat_dec (no associativity, at level 50). +Infix "==Z" := Z.eq_dec (no associativity, at level 50). + (* Definition const (A B : Type) (a : A) (b : B) : A := a. Definition compose (A B C : Type) (f : B -> C) (g : A -> B) (x : A) : C := f (g x). *) diff --git a/src/common/IntegerExtra.v b/src/common/IntegerExtra.v index 2f9aae6..5f06e26 100644 --- a/src/common/IntegerExtra.v +++ b/src/common/IntegerExtra.v @@ -51,6 +51,23 @@ Module PtrofsExtra. Ltac ptrofs_mod_tac m := repeat (ptrofs_mod_match m); lia. + Lemma signed_mod_unsigned_mod : + forall x m, + 0 < m -> + Ptrofs.modulus mod m = 0 -> + Ptrofs.signed x mod m = 0 -> + Ptrofs.unsigned x mod m = 0. + Proof. + intros. + + repeat match goal with + | [ _ : _ |- context[if ?x then _ else _] ] => destruct x + | [ _ : _ |- context[_ mod Ptrofs.modulus mod m] ] => + rewrite <- Zmod_div_mod; try lia; try assumption + | [ _ : _ |- context[Ptrofs.unsigned _] ] => rewrite Ptrofs.unsigned_signed + end; try (simplify; lia); ptrofs_mod_tac m. + Qed. + Lemma of_int_mod : forall x m, Int.signed x mod m = 0 -> @@ -121,6 +138,25 @@ Module PtrofsExtra. congruence. Qed. + Lemma divs_unsigned : + forall x y, + 0 <= Ptrofs.signed x -> + 0 < Ptrofs.signed y -> + Ptrofs.unsigned (Ptrofs.divs x y) = Ptrofs.signed x / Ptrofs.signed y. + Proof. + intros. + unfold Ptrofs.divs. + rewrite Ptrofs.unsigned_repr. + apply Z.quot_div_nonneg; lia. + split. + apply Z.quot_pos; lia. + apply Zquot.Zquot_le_upper_bound; try lia. + eapply Z.le_trans. + 2: { apply ZBinary.Z.le_mul_diag_r; simplify; try lia. } + pose proof (Ptrofs.signed_range x). + simplify; lia. + Qed. + Lemma mul_unsigned : forall x y, Ptrofs.mul x y = @@ -130,6 +166,20 @@ Module PtrofsExtra. apply Ptrofs.eqm_samerepr. apply Ptrofs.eqm_mult; exists 0; lia. Qed. + + Lemma pos_signed_unsigned : + forall x, + 0 <= Ptrofs.signed x -> + Ptrofs.signed x = Ptrofs.unsigned x. + Proof. + intros. + rewrite Ptrofs.unsigned_signed. + destruct (Ptrofs.lt x Ptrofs.zero) eqn:EQ. + unfold Ptrofs.lt in EQ. + destruct (zlt (Ptrofs.signed x) (Ptrofs.signed Ptrofs.zero)); try discriminate. + replace (Ptrofs.signed (Ptrofs.zero)) with 0 in l by reflexivity. lia. + reflexivity. + Qed. End PtrofsExtra. Module IntExtra. diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index cee04e9..1e16398 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -18,7 +18,7 @@ From compcert Require RTL Registers AST Integers. From compcert Require Import Globalenvs Memory. -From coqup Require Import Coquplib HTLgenspec HTLgen Value AssocMap Array IntegerExtra. +From coqup Require Import Coquplib HTLgenspec HTLgen Value AssocMap Array IntegerExtra ZExtra. From coqup Require HTL Verilog. Require Import Lia. @@ -196,6 +196,67 @@ Proof. assumption. Qed. +Lemma list_combine_lookup_first : + forall l1 l2 n, + length l1 = length l2 -> + nth_error l1 n = Some None -> + nth_error (list_combine Verilog.merge_cell l1 l2) n = nth_error l2 n. +Proof. + induction l1; intros; simplify. + + rewrite nth_error_nil in H0. + discriminate. + + destruct l2 eqn:EQl2. simplify. + simpl in H. invert H. + destruct n; simpl in *. + invert H0. simpl. reflexivity. + eauto. +Qed. + +Lemma combine_lookup_first : + forall a1 a2 n, + a1.(arr_length) = a2.(arr_length) -> + array_get_error n a1 = Some None -> + array_get_error n (combine Verilog.merge_cell a1 a2) = array_get_error n a2. +Proof. + intros. + + unfold array_get_error in *. + apply list_combine_lookup_first; eauto. + rewrite a1.(arr_wf). rewrite a2.(arr_wf). + assumption. +Qed. + +Lemma list_combine_lookup_second : + forall l1 l2 n x, + length l1 = length l2 -> + nth_error l1 n = Some (Some x) -> + nth_error (list_combine Verilog.merge_cell l1 l2) n = Some (Some x). +Proof. + induction l1; intros; simplify; auto. + + destruct l2 eqn:EQl2. simplify. + simpl in H. invert H. + destruct n; simpl in *. + invert H0. simpl. reflexivity. + eauto. +Qed. + +Lemma combine_lookup_second : + forall a1 a2 n x, + a1.(arr_length) = a2.(arr_length) -> + array_get_error n a1 = Some (Some x) -> + array_get_error n (combine Verilog.merge_cell a1 a2) = Some (Some x). +Proof. + intros. + + unfold array_get_error in *. + apply list_combine_lookup_second; eauto. + rewrite a1.(arr_wf). rewrite a2.(arr_wf). + assumption. +Qed. + (* Need to eventually move posToValue 32 to posToValueAuto, as that has this proof. *) Lemma assumption_32bit : forall v, @@ -1109,7 +1170,243 @@ Section CORRECTNESS. rewrite Registers.Regmap.gso; auto. - destruct c, chunk, addr, args; simplify; rt; simplify. - + admit. + + (** Preamble *) + invert MARR. simplify. + + unfold Op.eval_addressing in H0. + destruct (Archi.ptr64) eqn:ARCHI; simplify. + + unfold reg_stack_based_pointers in RSBP. + pose proof (RSBP r0) as RSBPr0. + + destruct (Registers.Regmap.get r0 rs) eqn:EQr0; simplify. + + rewrite ARCHI in H1. simplify. + subst. + + pose proof MASSOC as MASSOC'. + invert MASSOC'. + pose proof (H0 r0). + assert (HPler0 : Ple r0 (RTL.max_reg_function f)) + by (eapply RTL.max_reg_function_use; eauto; simplify; eauto). + apply H6 in HPler0. + invert HPler0; try congruence. + rewrite EQr0 in H8. + invert H8. + clear H0. clear H6. + + unfold check_address_parameter in *. simplify. + + remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0)) + (Integers.Ptrofs.of_int (Integers.Int.repr z))) as OFFSET. + + (** Read bounds assumption *) + assert (0 <= Integers.Ptrofs.signed OFFSET) as WRITE_BOUND_LOW by admit. + assert (Integers.Ptrofs.signed OFFSET < f.(RTL.fn_stacksize)) as WRITE_BOUND_HIGH by admit. + + (** Modular preservation proof *) + assert (Integers.Ptrofs.signed OFFSET mod 4 = 0) as MOD_PRESERVE. + { rewrite HeqOFFSET. + apply PtrofsExtra.add_mod; simplify; try lia. + exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *) + rewrite Integers.Ptrofs.signed_repr; try assumption. + admit. (* FIXME: Register bounds. *) + apply PtrofsExtra.of_int_mod. + rewrite Integers.Int.signed_repr; simplify; try split; try assumption. + } + + (* (** Normalisation proof *) *) + (* assert (Integers.Ptrofs.repr *) + (* (4 * Integers.Ptrofs.unsigned *) + (* (Integers.Ptrofs.divs OFFSET (Integers.Ptrofs.repr 4))) = OFFSET) *) + (* as NORMALISE. *) + (* { replace 4 with (Integers.Ptrofs.unsigned (Integers.Ptrofs.repr 4)) at 1 by reflexivity. *) + (* rewrite <- PtrofsExtra.mul_unsigned. *) + (* apply PtrofsExtra.mul_divs; auto. *) + (* rewrite Integers.Ptrofs.signed_repr; simplify; lia. } *) + + (* (** Normalised bounds proof *) *) + (* assert (0 <= *) + (* Integers.Ptrofs.unsigned (Integers.Ptrofs.divs OFFSET (Integers.Ptrofs.repr 4)) *) + (* < (RTL.fn_stacksize f / 4)) *) + (* as NORMALISE_BOUND. *) + (* { split. *) + (* apply Integers.Ptrofs.unsigned_range_2. *) + (* assert (forall x y, Integers.Ptrofs.divs x y = Integers.Ptrofs.divs x y ) by reflexivity. *) + (* unfold Integers.Ptrofs.divs at 2 in H0. *) + (* rewrite H0. clear H0. *) + (* rewrite Integers.Ptrofs.unsigned_repr; simplify. *) + (* rewrite Integers.Ptrofs.signed_repr; simplify; try lia. *) + (* rewrite Z.quot_div_nonneg; try (rewrite <- HeqOFFSET; assumption); try lia. *) + (* apply Zmult_lt_reg_r with (p := 4); try lia. *) + (* repeat rewrite ZLib.div_mul_undo; try lia. *) + (* rewrite Integers.Ptrofs.signed_repr. *) + (* rewrite Z.quot_div_nonneg; try (rewrite <- HeqOFFSET; assumption); try lia. *) + (* split. *) + (* apply Z.div_pos; try (rewrite <- HeqOFFSET; assumption); try lia. *) + (* apply Z.div_le_upper_bound; simplify; try (rewrite <- HeqOFFSET; lia); try lia. *) + (* simplify; lia. } *) + + (* inversion NORMALISE_BOUND as [ NORMALISE_BOUND_LOW NORMALISE_BOUND_HIGH ]; *) + (* clear NORMALISE_BOUND. *) + + eexists. split. + eapply Smallstep.plus_one. + eapply HTL.step_module; eauto. + econstructor. econstructor. econstructor. + eapply Verilog.stmnt_runp_Vnonblock_arr. simplify. + econstructor. + eapply Verilog.erun_Vbinop with (EQ := ?[EQ7]). + eapply Verilog.erun_Vbinop with (EQ := ?[EQ8]). + econstructor. + econstructor. + econstructor. econstructor. econstructor. econstructor. + econstructor. econstructor. econstructor. econstructor. + + all: simplify. + + (** State Lookup *) + unfold Verilog.merge_regs. + simplify. + unfold_merge. + apply AssocMap.gss. + + (** Match states *) + rewrite assumption_32bit. + econstructor; eauto. + + (** Match assocmaps *) + unfold Verilog.merge_regs. simplify. unfold_merge. + apply regs_lessdef_add_greater. apply greater_than_max_func. + assumption. + + (** States well formed *) + unfold state_st_wf. inversion 1. simplify. + unfold Verilog.merge_regs. + unfold_merge. + apply AssocMap.gss. + + (** Match stacks *) + admit. + + (** Equality proof *) + assert (Z.to_nat + (Integers.Ptrofs.unsigned + (Integers.Ptrofs.divs + OFFSET + (Integers.Ptrofs.repr 4))) + = + valueToNat (vdivs (vplus asr # r0 (ZToValue 32 z) ?EQ8) (ZToValue 32 4) ?EQ7)) + as EXPR_OK by admit. + + assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity. + inversion MASSOC; revert HeqOFFSET; subst; clear MASSOC; intros HeqOFFSET. + + econstructor. + repeat split; simplify. + unfold HTL.empty_stack. + simplify. + unfold Verilog.merge_arrs. + + rewrite AssocMap.gcombine. + 2: { reflexivity. } + unfold Verilog.arr_assocmap_set. + rewrite AssocMap.gss. + unfold Verilog.merge_arr. + rewrite AssocMap.gss. + setoid_rewrite H5. + reflexivity. + + rewrite combine_length. + rewrite <- array_set_len. + unfold arr_repeat. simplify. + apply list_repeat_len. + + rewrite <- array_set_len. + unfold arr_repeat. simplify. + rewrite list_repeat_len. + rewrite H4. reflexivity. + + intros. + destruct (4 * ptr ==Z Integers.Ptrofs.unsigned OFFSET). + + erewrite Mem.load_store_same. + 2: { rewrite ZERO. + rewrite Integers.Ptrofs.add_zero_l. + rewrite e. + rewrite Integers.Ptrofs.unsigned_repr. + exact H1. + apply Integers.Ptrofs.unsigned_range_2. } + constructor. + erewrite combine_lookup_second. + simpl. + assert (Ple src (RTL.max_reg_function f)) by admit. + apply H0 in H14. + destruct (Registers.Regmap.get src rs) eqn:EQ_SRC; constructor; invert H14; eauto. + + rewrite <- array_set_len. + unfold arr_repeat. simplify. + rewrite list_repeat_len. auto. + + assert (Z.to_nat ptr + = + valueToNat (vdivs (vplus asr # r0 (ZToValue 32 z) ?EQ8) (ZToValue 32 4) ?EQ7)) + as EXPR_OK' by admit. + + rewrite <- EXPR_OK'. + rewrite array_get_error_set_bound. + reflexivity. + unfold arr_length, arr_repeat. simpl. + rewrite list_repeat_len. lia. + + erewrite Mem.load_store_other with (m1 := m). + 2: { exact H1. } + 2: { right. + rewrite ZERO. + rewrite Integers.Ptrofs.add_zero_l. + rewrite Integers.Ptrofs.unsigned_repr. + simpl. + destruct (Z_le_gt_dec (4 * ptr + 4) (Integers.Ptrofs.unsigned OFFSET)); eauto. + right. + apply ZExtra.mod_0_bounds; try lia. + apply ZLib.Z_mod_mult'. + apply PtrofsExtra.signed_mod_unsigned_mod; eauto; lia. + split; try lia. + invert H13. + rewrite Z2Nat.id in H19; try lia. + apply Zmult_lt_compat_r with (p := 4) in H19; try lia. + rewrite ZLib.div_mul_undo in H19; try lia. + apply Z.le_trans with (m := RTL.fn_stacksize f); simplify; lia. + } + + rewrite <- EXPR_OK. + rewrite PtrofsExtra.divs_unsigned; try assumption. + destruct (ptr ==Z Integers.Ptrofs.unsigned OFFSET / 4). + apply Z.mul_cancel_r with (p := 4) in e; try lia. + rewrite <- PtrofsExtra.pos_signed_unsigned in e; try lia. + rewrite ZLib.div_mul_undo in e; try lia. + rewrite <- PtrofsExtra.pos_signed_unsigned in n; lia. + rewrite combine_lookup_first. + eapply H7; eauto. + + rewrite <- array_set_len. + unfold arr_repeat. simplify. + rewrite list_repeat_len. auto. + rewrite array_gso. + unfold array_get_error. + unfold arr_repeat. + simplify. + apply list_repeat_lookup. + lia. + replace (Integers.Ptrofs.signed (Integers.Ptrofs.repr 4)) with 4 by reflexivity. + rewrite <- PtrofsExtra.pos_signed_unsigned in n0; try lia. + intro. + apply Z2Nat.inj_iff in H14; try lia. + apply Z.div_pos; lia. + replace (Integers.Ptrofs.signed (Integers.Ptrofs.repr 4)) with 4 by reflexivity; lia. + + admit. + + admit. + admit. diff --git a/src/verilog/Array.v b/src/verilog/Array.v index fc52f04..f3e1cd7 100644 --- a/src/verilog/Array.v +++ b/src/verilog/Array.v @@ -41,6 +41,15 @@ Proof. Qed. Hint Resolve list_set_spec2 : array. +Lemma list_set_spec3 {A : Type} : + forall l i1 i2 (x : A), + i1 <> i2 -> + nth_error (list_set i1 x l) i2 = nth_error l i2. +Proof. + induction l; intros; destruct i1; destruct i2; simplify; try lia; try reflexivity; firstorder. +Qed. +Hint Resolve list_set_spec3 : array. + Lemma array_set_wf {A : Type} : forall l ln i (x : A), length l = ln -> length (list_set i x l) = ln. @@ -80,6 +89,13 @@ Proof. Qed. Hint Resolve array_set_spec2 : array. +Lemma array_set_len {A : Type} : + forall a i (x : A), + a.(arr_length) = (array_set i x a).(arr_length). +Proof. + unfold array_set. simplify. reflexivity. +Qed. + Definition array_get_error {A : Type} (i : nat) (a : Array A) : option A := nth_error a.(arr_contents) i. @@ -117,6 +133,19 @@ Proof. eauto with array. Qed. +Lemma array_gso {A : Type} : + forall (a : Array A) i1 i2 x, + i1 <> i2 -> + array_get_error i2 (array_set i1 x a) = array_get_error i2 a. +Proof. + intros. + + unfold array_get_error. + unfold array_set. + simplify. + eauto with array. +Qed. + Definition array_get {A : Type} (i : nat) (x : A) (a : Array A) : A := nth i a.(arr_contents) x. @@ -130,6 +159,17 @@ Proof. info_eauto with array. Qed. +Lemma array_get_get_error {A : Type} : + forall (a : Array A) i x d, + array_get_error i a = Some x -> + array_get i d a = x. +Proof. + intros. + unfold array_get. + unfold array_get_error in H. + auto using nth_error_nth. +Qed. + (** Tail recursive version of standard library function. *) Fixpoint list_repeat' {A : Type} (acc : list A) (a : A) (n : nat) : list A := match n with @@ -231,6 +271,19 @@ Proof. apply list_repeat'_cons. Qed. +Lemma list_repeat_lookup {A : Type} : + forall n i (a : A), + i < n -> + nth_error (list_repeat a n) i = Some a. +Proof. + induction n; intros. + + destruct i; simplify; lia. + + rewrite list_repeat_cons. + destruct i; simplify; firstorder. +Qed. + Definition arr_repeat {A : Type} (a : A) (n : nat) : Array A := make_array (list_repeat a n). Fixpoint list_combine {A B C : Type} (f : A -> B -> C) (x : list A) (y : list B) : list C := -- cgit From accf4b273525412801dc21c893d41c890c9fed6d Mon Sep 17 00:00:00 2001 From: James Pollard Date: Sun, 28 Jun 2020 17:15:14 +0100 Subject: Fix unsigned/signed issues. --- src/common/Coquplib.v | 23 ++++++--- src/common/IntegerExtra.v | 28 ++++++----- src/translation/HTLgen.v | 28 ++++++----- src/translation/HTLgenproof.v | 110 ++++++++++++++++++------------------------ 4 files changed, 92 insertions(+), 97 deletions(-) diff --git a/src/common/Coquplib.v b/src/common/Coquplib.v index b8a02d2..5de1e7c 100644 --- a/src/common/Coquplib.v +++ b/src/common/Coquplib.v @@ -69,45 +69,52 @@ Ltac kill_bools := Ltac unfold_constants := repeat match goal with - | [ _ : _ |- context[Integers.Ptrofs.modulus] ] => + | [ |- context[Integers.Ptrofs.modulus] ] => replace Integers.Ptrofs.modulus with 4294967296 by reflexivity | [ H : context[Integers.Ptrofs.modulus] |- _ ] => replace Integers.Ptrofs.modulus with 4294967296 in H by reflexivity - | [ _ : _ |- context[Integers.Ptrofs.min_signed] ] => + | [ |- context[Integers.Ptrofs.min_signed] ] => replace Integers.Ptrofs.min_signed with (-2147483648) by reflexivity | [ H : context[Integers.Ptrofs.min_signed] |- _ ] => replace Integers.Ptrofs.min_signed with (-2147483648) in H by reflexivity - | [ _ : _ |- context[Integers.Ptrofs.max_signed] ] => + | [ |- context[Integers.Ptrofs.max_signed] ] => replace Integers.Ptrofs.max_signed with 2147483647 by reflexivity | [ H : context[Integers.Ptrofs.max_signed] |- _ ] => replace Integers.Ptrofs.max_signed with 2147483647 in H by reflexivity - | [ _ : _ |- context[Integers.Ptrofs.max_unsigned] ] => + | [ |- context[Integers.Ptrofs.max_unsigned] ] => replace Integers.Ptrofs.max_unsigned with 4294967295 by reflexivity | [ H : context[Integers.Ptrofs.max_unsigned] |- _ ] => replace Integers.Ptrofs.max_unsigned with 4294967295 in H by reflexivity - | [ _ : _ |- context[Integers.Int.modulus] ] => + | [ |- context[Integers.Int.modulus] ] => replace Integers.Int.modulus with 4294967296 by reflexivity | [ H : context[Integers.Int.modulus] |- _ ] => replace Integers.Int.modulus with 4294967296 in H by reflexivity - | [ _ : _ |- context[Integers.Int.min_signed] ] => + | [ |- context[Integers.Int.min_signed] ] => replace Integers.Int.min_signed with (-2147483648) by reflexivity | [ H : context[Integers.Int.min_signed] |- _ ] => replace Integers.Int.min_signed with (-2147483648) in H by reflexivity - | [ _ : _ |- context[Integers.Int.max_signed] ] => + | [ |- context[Integers.Int.max_signed] ] => replace Integers.Int.max_signed with 2147483647 by reflexivity | [ H : context[Integers.Int.max_signed] |- _ ] => replace Integers.Int.max_signed with 2147483647 in H by reflexivity - | [ _ : _ |- context[Integers.Int.max_unsigned] ] => + | [ |- context[Integers.Int.max_unsigned] ] => replace Integers.Int.max_unsigned with 4294967295 by reflexivity | [ H : context[Integers.Int.max_unsigned] |- _ ] => replace Integers.Int.max_unsigned with 4294967295 in H by reflexivity + + | [ |- context[Integers.Ptrofs.unsigned (Integers.Ptrofs.repr ?x) ] ] => + match (eval compute in (0 <=? x)) with + | true => replace (Integers.Ptrofs.unsigned (Integers.Ptrofs.repr x)) + with x by reflexivity + | false => idtac + end end. Ltac simplify := unfold_constants; simpl in *; diff --git a/src/common/IntegerExtra.v b/src/common/IntegerExtra.v index 5f06e26..ec1fb07 100644 --- a/src/common/IntegerExtra.v +++ b/src/common/IntegerExtra.v @@ -105,10 +105,10 @@ Module PtrofsExtra. (m | Ptrofs.modulus) -> Ptrofs.signed x mod m = 0 -> Ptrofs.signed y mod m = 0 -> - (Ptrofs.signed (Ptrofs.add x y)) mod m = 0. + (Ptrofs.unsigned (Ptrofs.add x y)) mod m = 0. Proof. intros. unfold Ptrofs.add. - rewrite Ptrofs.signed_repr_eq. + rewrite Ptrofs.unsigned_repr_eq. repeat match goal with | [ _ : _ |- context[if ?x then _ else _] ] => destruct x @@ -118,21 +118,23 @@ Module PtrofsExtra. end; try (simplify; lia); ptrofs_mod_tac m. Qed. - Lemma mul_divs : + Lemma mul_divu : forall x y, - 0 <= Ptrofs.signed y -> - 0 < Ptrofs.signed x -> - Ptrofs.signed y mod Ptrofs.signed x = 0 -> - (Integers.Ptrofs.mul x (Integers.Ptrofs.divs y x)) = y. + 0 < Ptrofs.unsigned x -> + Ptrofs.unsigned y mod Ptrofs.unsigned x = 0 -> + (Integers.Ptrofs.mul x (Integers.Ptrofs.divu y x)) = y. Proof. intros. - pose proof (Ptrofs.mods_divs_Euclid y x). - pose proof (Zquot.Zrem_Zmod_zero (Ptrofs.signed y) (Ptrofs.signed x)). - apply <- H3 in H1; try lia; clear H3. - unfold Ptrofs.mods in H2. - rewrite H1 in H2. - replace (Ptrofs.repr 0) with (Ptrofs.zero) in H2 by reflexivity. + assert (x <> Ptrofs.zero). + { intro. + rewrite H1 in H. + replace (Ptrofs.unsigned Ptrofs.zero) with 0 in H by reflexivity. + lia. } + + exploit (Ptrofs.modu_divu_Euclid y x); auto; intros. + unfold Ptrofs.modu in H2. rewrite H0 in H2. + replace (Ptrofs.repr 0) with Ptrofs.zero in H2 by reflexivity. rewrite Ptrofs.add_zero in H2. rewrite Ptrofs.mul_commut. congruence. diff --git a/src/translation/HTLgen.v b/src/translation/HTLgen.v index 54ad77a..59fb70a 100644 --- a/src/translation/HTLgen.v +++ b/src/translation/HTLgen.v @@ -280,32 +280,36 @@ Definition translate_condition (c : Op.condition) (args : list reg) : mon expr : | _, _ => error (Errors.msg "Htlgen: condition instruction not implemented: other") end. -Definition check_address_parameter (p : Z) : bool := +Definition check_address_parameter_signed (p : Z) : bool := Z.eqb (Z.modulo p 4) 0 && Z.leb Integers.Ptrofs.min_signed p && Z.leb p Integers.Ptrofs.max_signed. +Definition check_address_parameter_unsigned (p : Z) : bool := + Z.eqb (Z.modulo p 4) 0 + && Z.leb p Integers.Ptrofs.max_unsigned. + Definition translate_eff_addressing (a: Op.addressing) (args: list reg) : mon expr := match a, args with (* TODO: We should be more methodical here; what are the possibilities?*) | Op.Aindexed off, r1::nil => - if (check_address_parameter off) + if (check_address_parameter_signed off) then ret (boplitz Vadd r1 off) else error (Errors.msg "Veriloggen: translate_eff_addressing address misaligned") | Op.Ascaled scale offset, r1::nil => - if (check_address_parameter scale) && (check_address_parameter offset) + if (check_address_parameter_signed scale) && (check_address_parameter_signed offset) then ret (Vbinop Vadd (boplitz Vmul r1 scale) (Vlit (ZToValue 32 offset))) else error (Errors.msg "Veriloggen: translate_eff_addressing address misaligned") | Op.Aindexed2 offset, r1::r2::nil => - if (check_address_parameter offset) + if (check_address_parameter_signed offset) then ret (Vbinop Vadd (Vvar r1) (boplitz Vadd r2 offset)) else error (Errors.msg "Veriloggen: translate_eff_addressing address misaligned") | Op.Aindexed2scaled scale offset, r1::r2::nil => (* Typical for dynamic array addressing *) - if (check_address_parameter scale) && (check_address_parameter offset) + if (check_address_parameter_signed scale) && (check_address_parameter_signed offset) then ret (Vbinop Vadd (boplitz Vadd r1 offset) (boplitz Vmul r2 scale)) else error (Errors.msg "Veriloggen: translate_eff_addressing address misaligned") | Op.Ainstack a, nil => (* We need to be sure that the base address is aligned *) let a := Integers.Ptrofs.unsigned a in - if (check_address_parameter a) + if (check_address_parameter_unsigned a) then ret (Vlit (ZToValue 32 a)) else error (Errors.msg "Veriloggen: translate_eff_addressing address misaligned") | _, _ => error (Errors.msg "Veriloggen: translate_eff_addressing unsuported addressing") @@ -392,19 +396,19 @@ Definition translate_arr_access (mem : AST.memory_chunk) (addr : Op.addressing) (args : list reg) (stack : reg) : mon expr := match mem, addr, args with (* TODO: We should be more methodical here; what are the possibilities?*) | Mint32, Op.Aindexed off, r1::nil => - if (check_address_parameter off) - then ret (Vvari stack (Vbinop Vdiv (boplitz Vadd r1 off) (Vlit (ZToValue 32 4)))) + if (check_address_parameter_signed off) + then ret (Vvari stack (Vbinop Vdivu (boplitz Vadd r1 off) (Vlit (ZToValue 32 4)))) else error (Errors.msg "HTLgen: translate_arr_access address misaligned") | Mint32, Op.Aindexed2scaled scale offset, r1::r2::nil => (* Typical for dynamic array addressing *) - if (check_address_parameter scale) && (check_address_parameter offset) + if (check_address_parameter_signed scale) && (check_address_parameter_signed offset) then ret (Vvari stack - (Vbinop Vdiv + (Vbinop Vdivu (Vbinop Vadd (boplitz Vadd r1 offset) (boplitz Vmul r2 scale)) (ZToValue 32 4))) else error (Errors.msg "HTLgen: translate_arr_access address misaligned") | Mint32, Op.Ainstack a, nil => (* We need to be sure that the base address is aligned *) - let a := Integers.Ptrofs.signed a in - if (check_address_parameter a) + let a := Integers.Ptrofs.unsigned a in + if (check_address_parameter_unsigned a) then ret (Vvari stack (Vlit (ZToValue 32 (a / 4)))) else error (Errors.msg "HTLgen: eff_addressing misaligned stack offset") | _, _, _ => error (Errors.msg "HTLgen: translate_arr_access unsuported addressing") diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index 1e16398..eb9ff7a 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -594,17 +594,17 @@ Section CORRECTNESS. invert H8. clear H0. clear H6. - unfold check_address_parameter in *. simplify. + unfold check_address_parameter_signed in *; + unfold check_address_parameter_unsigned in *; simplify. remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0)) (Integers.Ptrofs.of_int (Integers.Int.repr z))) as OFFSET. (** Read bounds assumption *) - assert (0 <= Integers.Ptrofs.signed OFFSET) as READ_BOUND_LOW by admit. - assert (Integers.Ptrofs.signed OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH by admit. + assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH by admit. (** Modular preservation proof *) - assert (Integers.Ptrofs.signed OFFSET mod 4 = 0) as MOD_PRESERVE. + assert (Integers.Ptrofs.unsigned OFFSET mod 4 = 0) as MOD_PRESERVE. { rewrite HeqOFFSET. apply PtrofsExtra.add_mod; simplify; try lia. exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *) @@ -617,34 +617,28 @@ Section CORRECTNESS. (** Normalisation proof *) assert (Integers.Ptrofs.repr (4 * Integers.Ptrofs.unsigned - (Integers.Ptrofs.divs OFFSET (Integers.Ptrofs.repr 4))) = OFFSET) + (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4))) = OFFSET) as NORMALISE. { replace 4 with (Integers.Ptrofs.unsigned (Integers.Ptrofs.repr 4)) at 1 by reflexivity. rewrite <- PtrofsExtra.mul_unsigned. - apply PtrofsExtra.mul_divs; auto. - rewrite Integers.Ptrofs.signed_repr; simplify; lia. } + apply PtrofsExtra.mul_divu; simplify; auto; lia. } (** Normalised bounds proof *) assert (0 <= - Integers.Ptrofs.unsigned (Integers.Ptrofs.divs OFFSET (Integers.Ptrofs.repr 4)) + Integers.Ptrofs.unsigned (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4)) < (RTL.fn_stacksize f / 4)) as NORMALISE_BOUND. { split. apply Integers.Ptrofs.unsigned_range_2. - assert (forall x y, Integers.Ptrofs.divs x y = Integers.Ptrofs.divs x y ) by reflexivity. - unfold Integers.Ptrofs.divs at 2 in H0. + assert (forall x y, Integers.Ptrofs.divu x y = Integers.Ptrofs.divu x y ) by reflexivity. + unfold Integers.Ptrofs.divu at 2 in H0. rewrite H0. clear H0. rewrite Integers.Ptrofs.unsigned_repr; simplify. - rewrite Integers.Ptrofs.signed_repr; simplify; try lia. - rewrite Z.quot_div_nonneg; try (rewrite <- HeqOFFSET; assumption); try lia. apply Zmult_lt_reg_r with (p := 4); try lia. repeat rewrite ZLib.div_mul_undo; try lia. - rewrite Integers.Ptrofs.signed_repr. - rewrite Z.quot_div_nonneg; try (rewrite <- HeqOFFSET; assumption); try lia. split. - apply Z.div_pos; try (rewrite <- HeqOFFSET; assumption); try lia. - apply Z.div_le_upper_bound; simplify; try (rewrite <- HeqOFFSET; lia); try lia. - simplify; lia. } + apply Z.div_pos; try lia; apply Integers.Ptrofs.unsigned_range_2. + apply Z.div_le_upper_bound; lia. } inversion NORMALISE_BOUND as [ NORMALISE_BOUND_LOW NORMALISE_BOUND_HIGH ]; clear NORMALISE_BOUND. @@ -687,7 +681,7 @@ Section CORRECTNESS. setoid_rewrite Integers.Ptrofs.add_zero_l in H7. specialize (H7 (Integers.Ptrofs.unsigned - (Integers.Ptrofs.divs + (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4)))). @@ -699,11 +693,11 @@ Section CORRECTNESS. assert (Z.to_nat (Integers.Ptrofs.unsigned - (Integers.Ptrofs.divs + (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4))) = - valueToNat (vdivs (vplus asr # r0 (ZToValue 32 z) ?EQ2) (ZToValue 32 4) ?EQ1)) + valueToNat (vdiv (vplus asr # r0 (ZToValue 32 z) ?EQ2) (ZToValue 32 4) ?EQ1)) as EXPR_OK by admit. rewrite <- EXPR_OK. rewrite NORMALISE in I. @@ -757,7 +751,7 @@ Section CORRECTNESS. rewrite Registers.Regmap.gss. unfold arr_stack_based_pointers in ASBP. specialize (ASBP (Integers.Ptrofs.unsigned - (Integers.Ptrofs.divs OFFSET (Integers.Ptrofs.repr 4)))). + (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4)))). exploit ASBP; auto; intros I. rewrite NORMALISE in I. @@ -804,7 +798,8 @@ Section CORRECTNESS. invert H10. invert H12. clear H0. clear H6. clear H8. - unfold check_address_parameter in *. simplify. + unfold check_address_parameter_signed in *; + unfold check_address_parameter_unsigned in *; simplify. remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0)) (Integers.Ptrofs.of_int @@ -812,11 +807,10 @@ Section CORRECTNESS. (Integers.Int.repr z0)))) as OFFSET. (** Read bounds assumption *) - assert (0 <= Integers.Ptrofs.signed OFFSET) as READ_BOUND_LOW by admit. - assert (Integers.Ptrofs.signed OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH by admit. + assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH by admit. (** Modular preservation proof *) - assert (Integers.Ptrofs.signed OFFSET mod 4 = 0) as MOD_PRESERVE. + assert (Integers.Ptrofs.unsigned OFFSET mod 4 = 0) as MOD_PRESERVE. { rewrite HeqOFFSET. apply PtrofsExtra.add_mod; simplify; try lia. exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *) @@ -835,34 +829,28 @@ Section CORRECTNESS. (** Normalisation proof *) assert (Integers.Ptrofs.repr (4 * Integers.Ptrofs.unsigned - (Integers.Ptrofs.divs OFFSET (Integers.Ptrofs.repr 4))) = OFFSET) + (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4))) = OFFSET) as NORMALISE. { replace 4 with (Integers.Ptrofs.unsigned (Integers.Ptrofs.repr 4)) at 1 by reflexivity. rewrite <- PtrofsExtra.mul_unsigned. - apply PtrofsExtra.mul_divs; auto. - rewrite Integers.Ptrofs.signed_repr; simplify; lia. } + apply PtrofsExtra.mul_divu; simplify; auto; lia. } (** Normalised bounds proof *) assert (0 <= - Integers.Ptrofs.unsigned (Integers.Ptrofs.divs OFFSET (Integers.Ptrofs.repr 4)) + Integers.Ptrofs.unsigned (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4)) < (RTL.fn_stacksize f / 4)) as NORMALISE_BOUND. { split. apply Integers.Ptrofs.unsigned_range_2. - assert (forall x y, Integers.Ptrofs.divs x y = Integers.Ptrofs.divs x y ) by reflexivity. - unfold Integers.Ptrofs.divs at 2 in H0. + assert (forall x y, Integers.Ptrofs.divu x y = Integers.Ptrofs.divu x y ) by reflexivity. + unfold Integers.Ptrofs.divu at 2 in H0. rewrite H0. clear H0. rewrite Integers.Ptrofs.unsigned_repr; simplify. - rewrite Integers.Ptrofs.signed_repr; simplify; try lia. - rewrite Z.quot_div_nonneg; try (rewrite <- HeqOFFSET; assumption); try lia. apply Zmult_lt_reg_r with (p := 4); try lia. repeat rewrite ZLib.div_mul_undo; try lia. - rewrite Integers.Ptrofs.signed_repr. - rewrite Z.quot_div_nonneg; try (rewrite <- HeqOFFSET; assumption); try lia. split. - apply Z.div_pos; try (rewrite <- HeqOFFSET; assumption); try lia. - apply Z.div_le_upper_bound; simplify; try (rewrite <- HeqOFFSET; lia); try lia. - simplify; lia. } + apply Z.div_pos; try lia; apply Integers.Ptrofs.unsigned_range_2. + apply Z.div_le_upper_bound; lia. } inversion NORMALISE_BOUND as [ NORMALISE_BOUND_LOW NORMALISE_BOUND_HIGH ]; clear NORMALISE_BOUND. @@ -911,7 +899,7 @@ Section CORRECTNESS. setoid_rewrite Integers.Ptrofs.add_zero_l in H7. specialize (H7 (Integers.Ptrofs.unsigned - (Integers.Ptrofs.divs + (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4)))). @@ -922,12 +910,12 @@ Section CORRECTNESS. intros I. assert (Z.to_nat (Integers.Ptrofs.unsigned - (Integers.Ptrofs.divs + (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4))) = valueToNat - (vdivs (vplus (vplus asr # r0 (ZToValue 32 z0) ?EQ5) - (vmul asr # r1 (ZToValue 32 z) ?EQ6) ?EQ4) (ZToValue 32 4) ?EQ3)) + (vdiv (vplus (vplus asr # r0 (ZToValue 32 z0) ?EQ5) + (vmul asr # r1 (ZToValue 32 z) ?EQ6) ?EQ4) (ZToValue 32 4) ?EQ3)) as EXPR_OK by admit. rewrite <- EXPR_OK. rewrite NORMALISE in I. @@ -981,7 +969,7 @@ Section CORRECTNESS. rewrite Registers.Regmap.gss. unfold arr_stack_based_pointers in ASBP. specialize (ASBP (Integers.Ptrofs.unsigned - (Integers.Ptrofs.divs OFFSET (Integers.Ptrofs.repr 4)))). + (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4)))). exploit ASBP; auto; intros I. rewrite NORMALISE in I. @@ -1001,7 +989,8 @@ Section CORRECTNESS. destruct (Archi.ptr64) eqn:ARCHI; simplify. rewrite ARCHI in H0. simplify. - unfold check_address_parameter in EQ; simplify. + unfold check_address_parameter_unsigned in *; + unfold check_address_parameter_signed in *; simplify. assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity. rewrite ZERO in H1. clear ZERO. @@ -1010,43 +999,36 @@ Section CORRECTNESS. remember i0 as OFFSET. (** Read bounds assumption *) - assert (0 <= Integers.Ptrofs.signed OFFSET) as READ_BOUND_LOW by admit. - assert (Integers.Ptrofs.signed OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH by admit. + assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH by admit. (** Modular preservation proof *) - rename H8 into MOD_PRESERVE. + rename H0 into MOD_PRESERVE. (** Normalisation proof *) assert (Integers.Ptrofs.repr (4 * Integers.Ptrofs.unsigned - (Integers.Ptrofs.divs OFFSET (Integers.Ptrofs.repr 4))) = OFFSET) + (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4))) = OFFSET) as NORMALISE. { replace 4 with (Integers.Ptrofs.unsigned (Integers.Ptrofs.repr 4)) at 1 by reflexivity. rewrite <- PtrofsExtra.mul_unsigned. - apply PtrofsExtra.mul_divs; auto. - rewrite Integers.Ptrofs.signed_repr; simplify; lia. } + apply PtrofsExtra.mul_divu; simplify; auto; try lia. } (** Normalised bounds proof *) assert (0 <= - Integers.Ptrofs.unsigned (Integers.Ptrofs.divs OFFSET (Integers.Ptrofs.repr 4)) + Integers.Ptrofs.unsigned (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4)) < (RTL.fn_stacksize f / 4)) as NORMALISE_BOUND. { split. apply Integers.Ptrofs.unsigned_range_2. - assert (forall x y, Integers.Ptrofs.divs x y = Integers.Ptrofs.divs x y ) by reflexivity. - unfold Integers.Ptrofs.divs at 2 in H0. + assert (forall x y, Integers.Ptrofs.divu x y = Integers.Ptrofs.divu x y ) by reflexivity. + unfold Integers.Ptrofs.divu at 2 in H0. rewrite H0. clear H0. rewrite Integers.Ptrofs.unsigned_repr; simplify. - rewrite Integers.Ptrofs.signed_repr; simplify; try lia. - rewrite Z.quot_div_nonneg; try (rewrite <- HeqOFFSET; assumption); try lia. apply Zmult_lt_reg_r with (p := 4); try lia. repeat rewrite ZLib.div_mul_undo; try lia. - rewrite Integers.Ptrofs.signed_repr. - rewrite Z.quot_div_nonneg; try (rewrite <- HeqOFFSET; assumption); try lia. split. - apply Z.div_pos; try (rewrite <- HeqOFFSET; assumption); try lia. - apply Z.div_le_upper_bound; simplify; try (rewrite <- HeqOFFSET; lia); try lia. - simplify; lia. } + apply Z.div_pos; try lia; apply Integers.Ptrofs.unsigned_range_2. + apply Z.div_le_upper_bound; lia. } inversion NORMALISE_BOUND as [ NORMALISE_BOUND_LOW NORMALISE_BOUND_HIGH ]; clear NORMALISE_BOUND. @@ -1086,7 +1068,7 @@ Section CORRECTNESS. setoid_rewrite Integers.Ptrofs.add_zero_l in H7. specialize (H7 (Integers.Ptrofs.unsigned - (Integers.Ptrofs.divs + (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4)))). @@ -1097,11 +1079,11 @@ Section CORRECTNESS. intros I. assert (Z.to_nat (Integers.Ptrofs.unsigned - (Integers.Ptrofs.divs + (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4))) = - valueToNat (ZToValue 32 (Integers.Ptrofs.signed OFFSET / 4))) + valueToNat (ZToValue 32 (Integers.Ptrofs.unsigned OFFSET / 4))) as EXPR_OK by admit. rewrite <- EXPR_OK. rewrite NORMALISE in I. @@ -1155,7 +1137,7 @@ Section CORRECTNESS. rewrite Registers.Regmap.gss. unfold arr_stack_based_pointers in ASBP. specialize (ASBP (Integers.Ptrofs.unsigned - (Integers.Ptrofs.divs OFFSET (Integers.Ptrofs.repr 4)))). + (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4)))). exploit ASBP; auto; intros I. rewrite NORMALISE in I. -- cgit From e9b4b89bd491fa91640ef56ccdacc6ecccc03908 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Sun, 28 Jun 2020 22:02:44 +0100 Subject: Finish first IStore proof (modulo some admissions). --- src/common/IntegerExtra.v | 24 +-- src/translation/HTLgenproof.v | 350 +++++++++++++++++++++++++++++++++++------- 2 files changed, 305 insertions(+), 69 deletions(-) diff --git a/src/common/IntegerExtra.v b/src/common/IntegerExtra.v index ec1fb07..8df70d9 100644 --- a/src/common/IntegerExtra.v +++ b/src/common/IntegerExtra.v @@ -140,23 +140,23 @@ Module PtrofsExtra. congruence. Qed. - Lemma divs_unsigned : + Lemma divu_unsigned : forall x y, - 0 <= Ptrofs.signed x -> - 0 < Ptrofs.signed y -> - Ptrofs.unsigned (Ptrofs.divs x y) = Ptrofs.signed x / Ptrofs.signed y. + 0 < Ptrofs.unsigned y -> + Ptrofs.unsigned x < Ptrofs.max_unsigned -> + Ptrofs.unsigned (Ptrofs.divu x y) = Ptrofs.unsigned x / Ptrofs.unsigned y. Proof. intros. - unfold Ptrofs.divs. - rewrite Ptrofs.unsigned_repr. - apply Z.quot_div_nonneg; lia. + unfold Ptrofs.divu. + rewrite Ptrofs.unsigned_repr; auto. split. - apply Z.quot_pos; lia. - apply Zquot.Zquot_le_upper_bound; try lia. + apply Z.div_pos; auto. + apply Ptrofs.unsigned_range. + apply Z.div_le_upper_bound; auto. eapply Z.le_trans. - 2: { apply ZBinary.Z.le_mul_diag_r; simplify; try lia. } - pose proof (Ptrofs.signed_range x). - simplify; lia. + apply Z.lt_le_incl. exact H0. + rewrite Z.mul_comm. + apply Z.le_mul_diag_r; simplify; lia. Qed. Lemma mul_unsigned : diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index eb9ff7a..ce92264 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -1177,17 +1177,17 @@ Section CORRECTNESS. invert H8. clear H0. clear H6. - unfold check_address_parameter in *. simplify. + unfold check_address_parameter_unsigned in *; + unfold check_address_parameter_signed in *; simplify. remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0)) (Integers.Ptrofs.of_int (Integers.Int.repr z))) as OFFSET. (** Read bounds assumption *) - assert (0 <= Integers.Ptrofs.signed OFFSET) as WRITE_BOUND_LOW by admit. - assert (Integers.Ptrofs.signed OFFSET < f.(RTL.fn_stacksize)) as WRITE_BOUND_HIGH by admit. + assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as WRITE_BOUND_HIGH by admit. (** Modular preservation proof *) - assert (Integers.Ptrofs.signed OFFSET mod 4 = 0) as MOD_PRESERVE. + assert (Integers.Ptrofs.unsigned OFFSET mod 4 = 0) as MOD_PRESERVE. { rewrite HeqOFFSET. apply PtrofsExtra.add_mod; simplify; try lia. exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *) @@ -1197,41 +1197,6 @@ Section CORRECTNESS. rewrite Integers.Int.signed_repr; simplify; try split; try assumption. } - (* (** Normalisation proof *) *) - (* assert (Integers.Ptrofs.repr *) - (* (4 * Integers.Ptrofs.unsigned *) - (* (Integers.Ptrofs.divs OFFSET (Integers.Ptrofs.repr 4))) = OFFSET) *) - (* as NORMALISE. *) - (* { replace 4 with (Integers.Ptrofs.unsigned (Integers.Ptrofs.repr 4)) at 1 by reflexivity. *) - (* rewrite <- PtrofsExtra.mul_unsigned. *) - (* apply PtrofsExtra.mul_divs; auto. *) - (* rewrite Integers.Ptrofs.signed_repr; simplify; lia. } *) - - (* (** Normalised bounds proof *) *) - (* assert (0 <= *) - (* Integers.Ptrofs.unsigned (Integers.Ptrofs.divs OFFSET (Integers.Ptrofs.repr 4)) *) - (* < (RTL.fn_stacksize f / 4)) *) - (* as NORMALISE_BOUND. *) - (* { split. *) - (* apply Integers.Ptrofs.unsigned_range_2. *) - (* assert (forall x y, Integers.Ptrofs.divs x y = Integers.Ptrofs.divs x y ) by reflexivity. *) - (* unfold Integers.Ptrofs.divs at 2 in H0. *) - (* rewrite H0. clear H0. *) - (* rewrite Integers.Ptrofs.unsigned_repr; simplify. *) - (* rewrite Integers.Ptrofs.signed_repr; simplify; try lia. *) - (* rewrite Z.quot_div_nonneg; try (rewrite <- HeqOFFSET; assumption); try lia. *) - (* apply Zmult_lt_reg_r with (p := 4); try lia. *) - (* repeat rewrite ZLib.div_mul_undo; try lia. *) - (* rewrite Integers.Ptrofs.signed_repr. *) - (* rewrite Z.quot_div_nonneg; try (rewrite <- HeqOFFSET; assumption); try lia. *) - (* split. *) - (* apply Z.div_pos; try (rewrite <- HeqOFFSET; assumption); try lia. *) - (* apply Z.div_le_upper_bound; simplify; try (rewrite <- HeqOFFSET; lia); try lia. *) - (* simplify; lia. } *) - - (* inversion NORMALISE_BOUND as [ NORMALISE_BOUND_LOW NORMALISE_BOUND_HIGH ]; *) - (* clear NORMALISE_BOUND. *) - eexists. split. eapply Smallstep.plus_one. eapply HTL.step_module; eauto. @@ -1274,11 +1239,11 @@ Section CORRECTNESS. (** Equality proof *) assert (Z.to_nat (Integers.Ptrofs.unsigned - (Integers.Ptrofs.divs + (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4))) = - valueToNat (vdivs (vplus asr # r0 (ZToValue 32 z) ?EQ8) (ZToValue 32 4) ?EQ7)) + valueToNat (vdiv (vplus asr # r0 (ZToValue 32 z) ?EQ8) (ZToValue 32 4) ?EQ7)) as EXPR_OK by admit. assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity. @@ -1322,7 +1287,8 @@ Section CORRECTNESS. constructor. erewrite combine_lookup_second. simpl. - assert (Ple src (RTL.max_reg_function f)) by admit. + assert (Ple src (RTL.max_reg_function f)) + by (eapply RTL.max_reg_function_use; eauto; simpl; auto); apply H0 in H14. destruct (Registers.Regmap.get src rs) eqn:EQ_SRC; constructor; invert H14; eauto. @@ -1330,12 +1296,248 @@ Section CORRECTNESS. unfold arr_repeat. simplify. rewrite list_repeat_len. auto. - assert (Z.to_nat ptr + assert (4 * ptr / 4 = Integers.Ptrofs.unsigned OFFSET / 4) by (f_equal; assumption). + rewrite Z.mul_comm in H14. + rewrite Z_div_mult in H14; try lia. + replace 4 with (Integers.Ptrofs.unsigned (Integers.Ptrofs.repr 4)) in H14 by reflexivity. + rewrite <- PtrofsExtra.divu_unsigned in H14; unfold_constants; try lia. + rewrite H14. rewrite EXPR_OK. + rewrite array_get_error_set_bound. + reflexivity. + unfold arr_length, arr_repeat. simpl. + rewrite list_repeat_len. lia. + + erewrite Mem.load_store_other with (m1 := m). + 2: { exact H1. } + 2: { right. + rewrite ZERO. + rewrite Integers.Ptrofs.add_zero_l. + rewrite Integers.Ptrofs.unsigned_repr. + simpl. + destruct (Z_le_gt_dec (4 * ptr + 4) (Integers.Ptrofs.unsigned OFFSET)); eauto. + right. + apply ZExtra.mod_0_bounds; try lia. + apply ZLib.Z_mod_mult'. + invert H13. + rewrite Z2Nat.id in H19; try lia. + apply Zmult_lt_compat_r with (p := 4) in H19; try lia. + rewrite ZLib.div_mul_undo in H19; try lia. + split; try lia. + apply Z.le_trans with (m := RTL.fn_stacksize f); simplify; lia. + } + + rewrite <- EXPR_OK. + rewrite PtrofsExtra.divu_unsigned; auto; try (unfold_constants; lia). + destruct (ptr ==Z Integers.Ptrofs.unsigned OFFSET / 4). + apply Z.mul_cancel_r with (p := 4) in e; try lia. + rewrite ZLib.div_mul_undo in e; try lia. + rewrite combine_lookup_first. + eapply H7; eauto. + + rewrite <- array_set_len. + unfold arr_repeat. simplify. + rewrite list_repeat_len. auto. + rewrite array_gso. + unfold array_get_error. + unfold arr_repeat. + simplify. + apply list_repeat_lookup. + lia. + unfold_constants. + intro. + apply Z2Nat.inj_iff in H14; try lia. + apply Z.div_pos; try lia. + apply Integers.Ptrofs.unsigned_range. + + assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity. + unfold arr_stack_based_pointers. + intros. + destruct (4 * ptr ==Z Integers.Ptrofs.unsigned OFFSET). + + simplify. + erewrite Mem.load_store_same. + 2: { rewrite ZERO. + rewrite Integers.Ptrofs.add_zero_l. + rewrite e. + rewrite Integers.Ptrofs.unsigned_repr. + exact H1. + apply Integers.Ptrofs.unsigned_range_2. } + simplify. + destruct (Registers.Regmap.get src rs) eqn:EQ_SRC; try constructor. + destruct (Archi.ptr64); try discriminate. + pose proof (RSBP src). rewrite EQ_SRC in H0. + assumption. + + simpl. + erewrite Mem.load_store_other with (m1 := m). + 2: { exact H1. } + 2: { right. + rewrite ZERO. + rewrite Integers.Ptrofs.add_zero_l. + rewrite Integers.Ptrofs.unsigned_repr. + simpl. + destruct (Z_le_gt_dec (4 * ptr + 4) (Integers.Ptrofs.unsigned OFFSET)); eauto. + right. + apply ZExtra.mod_0_bounds; try lia. + apply ZLib.Z_mod_mult'. + invert H0. + apply Zmult_lt_compat_r with (p := 4) in H14; try lia. + rewrite ZLib.div_mul_undo in H14; try lia. + split; try lia. + apply Z.le_trans with (m := RTL.fn_stacksize f); simplify; lia. + } + apply ASBP; assumption. + + + (** Preamble *) + invert MARR. simplify. + + unfold Op.eval_addressing in H0. + destruct (Archi.ptr64) eqn:ARCHI; simplify. + + unfold reg_stack_based_pointers in RSBP. + pose proof (RSBP r0) as RSBPr0. + + destruct (Registers.Regmap.get r0 rs) eqn:EQr0; simplify. + + rewrite ARCHI in H1. simplify. + subst. + + pose proof MASSOC as MASSOC'. + invert MASSOC'. + pose proof (H0 r0). + assert (HPler0 : Ple r0 (RTL.max_reg_function f)) + by (eapply RTL.max_reg_function_use; eauto; simplify; eauto). + apply H6 in HPler0. + invert HPler0; try congruence. + rewrite EQr0 in H8. + invert H8. + clear H0. clear H6. + + unfold check_address_parameter_unsigned in *; + unfold check_address_parameter_signed in *; simplify. + + remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0)) + (Integers.Ptrofs.of_int (Integers.Int.repr z))) as OFFSET. + + (** Read bounds assumption *) + assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as WRITE_BOUND_HIGH by admit. + + (** Modular preservation proof *) + assert (Integers.Ptrofs.unsigned OFFSET mod 4 = 0) as MOD_PRESERVE. + { rewrite HeqOFFSET. + apply PtrofsExtra.add_mod; simplify; try lia. + exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *) + rewrite Integers.Ptrofs.signed_repr; try assumption. + admit. (* FIXME: Register bounds. *) + apply PtrofsExtra.of_int_mod. + rewrite Integers.Int.signed_repr; simplify; try split; try assumption. + } + + eexists. split. + eapply Smallstep.plus_one. + eapply HTL.step_module; eauto. + econstructor. econstructor. econstructor. + eapply Verilog.stmnt_runp_Vnonblock_arr. simplify. + econstructor. + eapply Verilog.erun_Vbinop with (EQ := ?[EQ7]). + eapply Verilog.erun_Vbinop with (EQ := ?[EQ8]). + econstructor. + econstructor. + econstructor. econstructor. econstructor. econstructor. + econstructor. econstructor. econstructor. econstructor. + + all: simplify. + + (** State Lookup *) + unfold Verilog.merge_regs. + simplify. + unfold_merge. + apply AssocMap.gss. + + (** Match states *) + rewrite assumption_32bit. + econstructor; eauto. + + (** Match assocmaps *) + unfold Verilog.merge_regs. simplify. unfold_merge. + apply regs_lessdef_add_greater. apply greater_than_max_func. + assumption. + + (** States well formed *) + unfold state_st_wf. inversion 1. simplify. + unfold Verilog.merge_regs. + unfold_merge. + apply AssocMap.gss. + + (** Match stacks *) + admit. + + (** Equality proof *) + assert (Z.to_nat + (Integers.Ptrofs.unsigned + (Integers.Ptrofs.divu + OFFSET + (Integers.Ptrofs.repr 4))) = - valueToNat (vdivs (vplus asr # r0 (ZToValue 32 z) ?EQ8) (ZToValue 32 4) ?EQ7)) - as EXPR_OK' by admit. + valueToNat (vdiv (vplus asr # r0 (ZToValue 32 z) ?EQ8) (ZToValue 32 4) ?EQ7)) + as EXPR_OK by admit. + + assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity. + inversion MASSOC; revert HeqOFFSET; subst; clear MASSOC; intros HeqOFFSET. + + econstructor. + repeat split; simplify. + unfold HTL.empty_stack. + simplify. + unfold Verilog.merge_arrs. + + rewrite AssocMap.gcombine. + 2: { reflexivity. } + unfold Verilog.arr_assocmap_set. + rewrite AssocMap.gss. + unfold Verilog.merge_arr. + rewrite AssocMap.gss. + setoid_rewrite H5. + reflexivity. - rewrite <- EXPR_OK'. + rewrite combine_length. + rewrite <- array_set_len. + unfold arr_repeat. simplify. + apply list_repeat_len. + + rewrite <- array_set_len. + unfold arr_repeat. simplify. + rewrite list_repeat_len. + rewrite H4. reflexivity. + + intros. + destruct (4 * ptr ==Z Integers.Ptrofs.unsigned OFFSET). + + erewrite Mem.load_store_same. + 2: { rewrite ZERO. + rewrite Integers.Ptrofs.add_zero_l. + rewrite e. + rewrite Integers.Ptrofs.unsigned_repr. + exact H1. + apply Integers.Ptrofs.unsigned_range_2. } + constructor. + erewrite combine_lookup_second. + simpl. + assert (Ple src (RTL.max_reg_function f)) + by (eapply RTL.max_reg_function_use; eauto; simpl; auto); + apply H0 in H14. + destruct (Registers.Regmap.get src rs) eqn:EQ_SRC; constructor; invert H14; eauto. + + rewrite <- array_set_len. + unfold arr_repeat. simplify. + rewrite list_repeat_len. auto. + + assert (4 * ptr / 4 = Integers.Ptrofs.unsigned OFFSET / 4) by (f_equal; assumption). + rewrite Z.mul_comm in H14. + rewrite Z_div_mult in H14; try lia. + replace 4 with (Integers.Ptrofs.unsigned (Integers.Ptrofs.repr 4)) in H14 by reflexivity. + rewrite <- PtrofsExtra.divu_unsigned in H14; unfold_constants; try lia. + rewrite H14. rewrite EXPR_OK. rewrite array_get_error_set_bound. reflexivity. unfold arr_length, arr_repeat. simpl. @@ -1352,22 +1554,19 @@ Section CORRECTNESS. right. apply ZExtra.mod_0_bounds; try lia. apply ZLib.Z_mod_mult'. - apply PtrofsExtra.signed_mod_unsigned_mod; eauto; lia. - split; try lia. invert H13. rewrite Z2Nat.id in H19; try lia. apply Zmult_lt_compat_r with (p := 4) in H19; try lia. rewrite ZLib.div_mul_undo in H19; try lia. + split; try lia. apply Z.le_trans with (m := RTL.fn_stacksize f); simplify; lia. } rewrite <- EXPR_OK. - rewrite PtrofsExtra.divs_unsigned; try assumption. + rewrite PtrofsExtra.divu_unsigned; auto; try (unfold_constants; lia). destruct (ptr ==Z Integers.Ptrofs.unsigned OFFSET / 4). apply Z.mul_cancel_r with (p := 4) in e; try lia. - rewrite <- PtrofsExtra.pos_signed_unsigned in e; try lia. rewrite ZLib.div_mul_undo in e; try lia. - rewrite <- PtrofsExtra.pos_signed_unsigned in n; lia. rewrite combine_lookup_first. eapply H7; eauto. @@ -1380,16 +1579,53 @@ Section CORRECTNESS. simplify. apply list_repeat_lookup. lia. - replace (Integers.Ptrofs.signed (Integers.Ptrofs.repr 4)) with 4 by reflexivity. - rewrite <- PtrofsExtra.pos_signed_unsigned in n0; try lia. + unfold_constants. intro. apply Z2Nat.inj_iff in H14; try lia. - apply Z.div_pos; lia. - replace (Integers.Ptrofs.signed (Integers.Ptrofs.repr 4)) with 4 by reflexivity; lia. + apply Z.div_pos; try lia. + apply Integers.Ptrofs.unsigned_range. + + assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity. + unfold arr_stack_based_pointers. + intros. + destruct (4 * ptr ==Z Integers.Ptrofs.unsigned OFFSET). + + simplify. + erewrite Mem.load_store_same. + 2: { rewrite ZERO. + rewrite Integers.Ptrofs.add_zero_l. + rewrite e. + rewrite Integers.Ptrofs.unsigned_repr. + exact H1. + apply Integers.Ptrofs.unsigned_range_2. } + simplify. + destruct (Registers.Regmap.get src rs) eqn:EQ_SRC; try constructor. + destruct (Archi.ptr64); try discriminate. + pose proof (RSBP src). rewrite EQ_SRC in H0. + assumption. + + simpl. + erewrite Mem.load_store_other with (m1 := m). + 2: { exact H1. } + 2: { right. + rewrite ZERO. + rewrite Integers.Ptrofs.add_zero_l. + rewrite Integers.Ptrofs.unsigned_repr. + simpl. + destruct (Z_le_gt_dec (4 * ptr + 4) (Integers.Ptrofs.unsigned OFFSET)); eauto. + right. + apply ZExtra.mod_0_bounds; try lia. + apply ZLib.Z_mod_mult'. + invert H0. + apply Zmult_lt_compat_r with (p := 4) in H14; try lia. + rewrite ZLib.div_mul_undo in H14; try lia. + split; try lia. + apply Z.le_trans with (m := RTL.fn_stacksize f); simplify; lia. + } + apply ASBP; assumption. + - admit. - + admit. + admit. - eexists. split. apply Smallstep.plus_one. -- cgit From b988e11ff068e6c27a5252ed39113ca2bebf35e8 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Sun, 28 Jun 2020 22:14:36 +0100 Subject: Fix second IStore proof. --- src/translation/HTLgenproof.v | 79 +++++++++++++++++++++++++++---------------- 1 file changed, 50 insertions(+), 29 deletions(-) diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index ce92264..37cebb6 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -1197,6 +1197,7 @@ Section CORRECTNESS. rewrite Integers.Int.signed_repr; simplify; try split; try assumption. } + (** Start of proof proper *) eexists. split. eapply Smallstep.plus_one. eapply HTL.step_module; eauto. @@ -1396,31 +1397,41 @@ Section CORRECTNESS. unfold reg_stack_based_pointers in RSBP. pose proof (RSBP r0) as RSBPr0. + pose proof (RSBP r1) as RSBPr1. - destruct (Registers.Regmap.get r0 rs) eqn:EQr0; simplify. + destruct (Registers.Regmap.get r0 rs) eqn:EQr0; + destruct (Registers.Regmap.get r1 rs) eqn:EQr1; simplify. rewrite ARCHI in H1. simplify. subst. + clear RSBPr1. pose proof MASSOC as MASSOC'. invert MASSOC'. pose proof (H0 r0). + pose proof (H0 r1). assert (HPler0 : Ple r0 (RTL.max_reg_function f)) by (eapply RTL.max_reg_function_use; eauto; simplify; eauto). + assert (HPler1 : Ple r1 (RTL.max_reg_function f)) + by (eapply RTL.max_reg_function_use; eauto; simpl; auto). apply H6 in HPler0. - invert HPler0; try congruence. - rewrite EQr0 in H8. - invert H8. - clear H0. clear H6. + apply H8 in HPler1. + invert HPler0; invert HPler1; try congruence. + rewrite EQr0 in H10. + rewrite EQr1 in H12. + invert H10. invert H12. + clear H0. clear H6. clear H8. - unfold check_address_parameter_unsigned in *; - unfold check_address_parameter_signed in *; simplify. + unfold check_address_parameter_signed in *; + unfold check_address_parameter_unsigned in *; simplify. remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0)) - (Integers.Ptrofs.of_int (Integers.Int.repr z))) as OFFSET. + (Integers.Ptrofs.of_int + (Integers.Int.add (Integers.Int.mul (valueToInt asr # r1) (Integers.Int.repr z)) + (Integers.Int.repr z0)))) as OFFSET. (** Read bounds assumption *) - assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as WRITE_BOUND_HIGH by admit. + assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH by admit. (** Modular preservation proof *) assert (Integers.Ptrofs.unsigned OFFSET mod 4 = 0) as MOD_PRESERVE. @@ -1430,19 +1441,29 @@ Section CORRECTNESS. rewrite Integers.Ptrofs.signed_repr; try assumption. admit. (* FIXME: Register bounds. *) apply PtrofsExtra.of_int_mod. + apply IntExtra.add_mod; simplify; try lia. + exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *) + apply IntExtra.mul_mod; simplify; try lia. + exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *) + admit. (* FIXME: Register bounds. *) + rewrite Integers.Int.signed_repr; simplify; try split; try assumption. rewrite Integers.Int.signed_repr; simplify; try split; try assumption. } + (** Start of proof proper *) eexists. split. eapply Smallstep.plus_one. eapply HTL.step_module; eauto. econstructor. econstructor. econstructor. eapply Verilog.stmnt_runp_Vnonblock_arr. simplify. econstructor. - eapply Verilog.erun_Vbinop with (EQ := ?[EQ7]). - eapply Verilog.erun_Vbinop with (EQ := ?[EQ8]). - econstructor. + eapply Verilog.erun_Vbinop with (EQ := ?[EQ9]). + eapply Verilog.erun_Vbinop with (EQ := ?[EQ10]). + eapply Verilog.erun_Vbinop with (EQ := ?[EQ11]). + econstructor. econstructor. econstructor. econstructor. econstructor. + eapply Verilog.erun_Vbinop with (EQ := ?[EQ12]). + econstructor. econstructor. econstructor. econstructor. econstructor. econstructor. econstructor. econstructor. econstructor. econstructor. econstructor. econstructor. @@ -1479,7 +1500,9 @@ Section CORRECTNESS. OFFSET (Integers.Ptrofs.repr 4))) = - valueToNat (vdiv (vplus asr # r0 (ZToValue 32 z) ?EQ8) (ZToValue 32 4) ?EQ7)) + valueToNat (vdiv + (vplus (vplus asr # r0 (ZToValue 32 z0) ?EQ11) (vmul asr # r1 (ZToValue 32 z) ?EQ12) + ?EQ10) (ZToValue 32 4) ?EQ9)) as EXPR_OK by admit. assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity. @@ -1525,19 +1548,19 @@ Section CORRECTNESS. simpl. assert (Ple src (RTL.max_reg_function f)) by (eapply RTL.max_reg_function_use; eauto; simpl; auto); - apply H0 in H14. - destruct (Registers.Regmap.get src rs) eqn:EQ_SRC; constructor; invert H14; eauto. + apply H0 in H21. + destruct (Registers.Regmap.get src rs) eqn:EQ_SRC; constructor; invert H21; eauto. rewrite <- array_set_len. unfold arr_repeat. simplify. rewrite list_repeat_len. auto. assert (4 * ptr / 4 = Integers.Ptrofs.unsigned OFFSET / 4) by (f_equal; assumption). - rewrite Z.mul_comm in H14. - rewrite Z_div_mult in H14; try lia. - replace 4 with (Integers.Ptrofs.unsigned (Integers.Ptrofs.repr 4)) in H14 by reflexivity. - rewrite <- PtrofsExtra.divu_unsigned in H14; unfold_constants; try lia. - rewrite H14. rewrite EXPR_OK. + rewrite Z.mul_comm in H21. + rewrite Z_div_mult in H21; try lia. + replace 4 with (Integers.Ptrofs.unsigned (Integers.Ptrofs.repr 4)) in H21 by reflexivity. + rewrite <- PtrofsExtra.divu_unsigned in H21; unfold_constants; try lia. + rewrite H21. rewrite EXPR_OK. rewrite array_get_error_set_bound. reflexivity. unfold arr_length, arr_repeat. simpl. @@ -1554,10 +1577,10 @@ Section CORRECTNESS. right. apply ZExtra.mod_0_bounds; try lia. apply ZLib.Z_mod_mult'. - invert H13. - rewrite Z2Nat.id in H19; try lia. - apply Zmult_lt_compat_r with (p := 4) in H19; try lia. - rewrite ZLib.div_mul_undo in H19; try lia. + invert H20. + rewrite Z2Nat.id in H22; try lia. + apply Zmult_lt_compat_r with (p := 4) in H22; try lia. + rewrite ZLib.div_mul_undo in H22; try lia. split; try lia. apply Z.le_trans with (m := RTL.fn_stacksize f); simplify; lia. } @@ -1581,7 +1604,7 @@ Section CORRECTNESS. lia. unfold_constants. intro. - apply Z2Nat.inj_iff in H14; try lia. + apply Z2Nat.inj_iff in H21; try lia. apply Z.div_pos; try lia. apply Integers.Ptrofs.unsigned_range. @@ -1617,15 +1640,13 @@ Section CORRECTNESS. apply ZExtra.mod_0_bounds; try lia. apply ZLib.Z_mod_mult'. invert H0. - apply Zmult_lt_compat_r with (p := 4) in H14; try lia. - rewrite ZLib.div_mul_undo in H14; try lia. + apply Zmult_lt_compat_r with (p := 4) in H21; try lia. + rewrite ZLib.div_mul_undo in H21; try lia. split; try lia. apply Z.le_trans with (m := RTL.fn_stacksize f); simplify; lia. } apply ASBP; assumption. - - + admit. - eexists. split. apply Smallstep.plus_one. -- cgit From 2f71ed762e496545699ba804e29c573aa2e0b947 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Sun, 28 Jun 2020 22:22:17 +0100 Subject: Finish store proof modulo: * EXPR_OK proofs (Yann). * Trivial register size proof (i.e. register values < 2^32). * Read bounds (to be extracted from RTL semantics). * Stack frame proof issues. --- src/translation/HTLgenproof.v | 207 +++++++++++++++++++++++++++++++++++++++++- 1 file changed, 206 insertions(+), 1 deletion(-) diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index 37cebb6..08b1216 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -1647,7 +1647,212 @@ Section CORRECTNESS. } apply ASBP; assumption. - + admit. + + invert MARR. simplify. + + unfold Op.eval_addressing in H0. + destruct (Archi.ptr64) eqn:ARCHI; simplify. + rewrite ARCHI in H0. simplify. + + unfold check_address_parameter_unsigned in *; + unfold check_address_parameter_signed in *; simplify. + + assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity. + rewrite ZERO in H1. clear ZERO. + rewrite Integers.Ptrofs.add_zero_l in H1. + + remember i0 as OFFSET. + + (** Read bounds assumption *) + assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH by admit. + + (** Modular preservation proof *) + rename H0 into MOD_PRESERVE. + + (** Start of proof proper *) + eexists. split. + eapply Smallstep.plus_one. + eapply HTL.step_module; eauto. + econstructor. econstructor. econstructor. + eapply Verilog.stmnt_runp_Vnonblock_arr. simplify. + econstructor. econstructor. econstructor. econstructor. + + all: simplify. + + (** State Lookup *) + unfold Verilog.merge_regs. + simplify. + unfold_merge. + apply AssocMap.gss. + + (** Match states *) + rewrite assumption_32bit. + econstructor; eauto. + + (** Match assocmaps *) + unfold Verilog.merge_regs. simplify. unfold_merge. + apply regs_lessdef_add_greater. apply greater_than_max_func. + assumption. + + (** States well formed *) + unfold state_st_wf. inversion 1. simplify. + unfold Verilog.merge_regs. + unfold_merge. + apply AssocMap.gss. + + (** Match stacks *) + admit. + + (** Equality proof *) + assert (Z.to_nat + (Integers.Ptrofs.unsigned + (Integers.Ptrofs.divu + OFFSET + (Integers.Ptrofs.repr 4))) + = + valueToNat (ZToValue 32 (Integers.Ptrofs.unsigned OFFSET / 4))) + as EXPR_OK by admit. + + assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity. + inversion MASSOC; revert HeqOFFSET; subst; clear MASSOC; intros HeqOFFSET. + + econstructor. + repeat split; simplify. + unfold HTL.empty_stack. + simplify. + unfold Verilog.merge_arrs. + + rewrite AssocMap.gcombine. + 2: { reflexivity. } + unfold Verilog.arr_assocmap_set. + rewrite AssocMap.gss. + unfold Verilog.merge_arr. + rewrite AssocMap.gss. + setoid_rewrite H5. + reflexivity. + + rewrite combine_length. + rewrite <- array_set_len. + unfold arr_repeat. simplify. + apply list_repeat_len. + + rewrite <- array_set_len. + unfold arr_repeat. simplify. + rewrite list_repeat_len. + rewrite H4. reflexivity. + + intros. + destruct (4 * ptr ==Z Integers.Ptrofs.unsigned OFFSET). + + erewrite Mem.load_store_same. + 2: { rewrite ZERO. + rewrite Integers.Ptrofs.add_zero_l. + rewrite e. + rewrite Integers.Ptrofs.unsigned_repr. + exact H1. + apply Integers.Ptrofs.unsigned_range_2. } + constructor. + erewrite combine_lookup_second. + simpl. + assert (Ple src (RTL.max_reg_function f)) + by (eapply RTL.max_reg_function_use; eauto; simpl; auto); + apply H0 in H10. + destruct (Registers.Regmap.get src rs) eqn:EQ_SRC; constructor; invert H10; eauto. + + rewrite <- array_set_len. + unfold arr_repeat. simplify. + rewrite list_repeat_len. auto. + + assert (4 * ptr / 4 = Integers.Ptrofs.unsigned OFFSET / 4) by (f_equal; assumption). + rewrite Z.mul_comm in H10. + rewrite Z_div_mult in H10; try lia. + replace 4 with (Integers.Ptrofs.unsigned (Integers.Ptrofs.repr 4)) in H10 by reflexivity. + rewrite <- PtrofsExtra.divu_unsigned in H10; unfold_constants; try lia. + rewrite H10. rewrite EXPR_OK. + rewrite array_get_error_set_bound. + reflexivity. + unfold arr_length, arr_repeat. simpl. + rewrite list_repeat_len. lia. + + erewrite Mem.load_store_other with (m1 := m). + 2: { exact H1. } + 2: { right. + rewrite ZERO. + rewrite Integers.Ptrofs.add_zero_l. + rewrite Integers.Ptrofs.unsigned_repr. + simpl. + destruct (Z_le_gt_dec (4 * ptr + 4) (Integers.Ptrofs.unsigned OFFSET)); eauto. + right. + apply ZExtra.mod_0_bounds; try lia. + apply ZLib.Z_mod_mult'. + invert H8. + rewrite Z2Nat.id in H12; try lia. + apply Zmult_lt_compat_r with (p := 4) in H12; try lia. + rewrite ZLib.div_mul_undo in H12; try lia. + split; try lia. + apply Z.le_trans with (m := RTL.fn_stacksize f); simplify; lia. + } + + rewrite <- EXPR_OK. + rewrite PtrofsExtra.divu_unsigned; auto; try (unfold_constants; lia). + destruct (ptr ==Z Integers.Ptrofs.unsigned OFFSET / 4). + apply Z.mul_cancel_r with (p := 4) in e; try lia. + rewrite ZLib.div_mul_undo in e; try lia. + rewrite combine_lookup_first. + eapply H7; eauto. + + rewrite <- array_set_len. + unfold arr_repeat. simplify. + rewrite list_repeat_len. auto. + rewrite array_gso. + unfold array_get_error. + unfold arr_repeat. + simplify. + apply list_repeat_lookup. + lia. + unfold_constants. + intro. + apply Z2Nat.inj_iff in H10; try lia. + apply Z.div_pos; try lia. + apply Integers.Ptrofs.unsigned_range. + + assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity. + unfold arr_stack_based_pointers. + intros. + destruct (4 * ptr ==Z Integers.Ptrofs.unsigned OFFSET). + + simplify. + erewrite Mem.load_store_same. + 2: { rewrite ZERO. + rewrite Integers.Ptrofs.add_zero_l. + rewrite e. + rewrite Integers.Ptrofs.unsigned_repr. + exact H1. + apply Integers.Ptrofs.unsigned_range_2. } + simplify. + destruct (Registers.Regmap.get src rs) eqn:EQ_SRC; try constructor. + destruct (Archi.ptr64); try discriminate. + pose proof (RSBP src). rewrite EQ_SRC in H0. + assumption. + + simpl. + erewrite Mem.load_store_other with (m1 := m). + 2: { exact H1. } + 2: { right. + rewrite ZERO. + rewrite Integers.Ptrofs.add_zero_l. + rewrite Integers.Ptrofs.unsigned_repr. + simpl. + destruct (Z_le_gt_dec (4 * ptr + 4) (Integers.Ptrofs.unsigned OFFSET)); eauto. + right. + apply ZExtra.mod_0_bounds; try lia. + apply ZLib.Z_mod_mult'. + invert H0. + apply Zmult_lt_compat_r with (p := 4) in H10; try lia. + rewrite ZLib.div_mul_undo in H10; try lia. + split; try lia. + apply Z.le_trans with (m := RTL.fn_stacksize f); simplify; lia. + } + apply ASBP; assumption. - eexists. split. apply Smallstep.plus_one. eapply HTL.step_module; eauto. -- cgit From b56f06b184afe0b1a735ac91cb450784f642d45e Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sun, 28 Jun 2020 23:23:32 +0100 Subject: Add yosys to dependencies --- Makefile | 2 +- shell.nix | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/Makefile b/Makefile index 0167a78..3d31b2f 100644 --- a/Makefile +++ b/Makefile @@ -28,7 +28,7 @@ all: lib/COMPCERTSTAMP $(MAKE) compile lib/COMPCERTSTAMP: - (cd lib/CompCert && ./configure $(ARCH)) + (cd lib/CompCert && ./configure --ignore-coq-version $(ARCH)) $(MAKE) -C lib/CompCert touch $@ diff --git a/shell.nix b/shell.nix index f81ded1..3906659 100644 --- a/shell.nix +++ b/shell.nix @@ -1,5 +1,5 @@ with import {}; mkShell { - buildInputs = (import ./.).buildInputs ++ [ocamlPackages.ocp-indent verilog]; + buildInputs = (import ./.).buildInputs ++ [ocamlPackages.ocp-indent verilog yosys]; } -- cgit From 0c360ec297c42d73c1090958d061447c2bfbe31b Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sun, 28 Jun 2020 23:48:06 +0100 Subject: Fix proof again with Verilog semantics changes --- src/translation/HTLgenproof.v | 12 +++++++++++- src/verilog/Value.v | 12 ++++++++++++ 2 files changed, 23 insertions(+), 1 deletion(-) diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index 753dccf..2f296f2 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -18,7 +18,7 @@ From compcert Require RTL Registers AST Integers. From compcert Require Import Globalenvs Memory. -From coqup Require Import Coquplib HTLgenspec HTLgen Value AssocMap Array IntegerExtra ZExtra. +From coqup Require Import Coquplib HTLgenspec HTLgen Value AssocMap Array IntegerExtra. From coqup Require HTL Verilog. Require Import Lia. @@ -427,6 +427,7 @@ Section CORRECTNESS. split. apply Smallstep.plus_one. eapply HTL.step_module; eauto. + apply assumption_32bit. (* processing of state *) econstructor. simplify. @@ -646,6 +647,7 @@ Section CORRECTNESS. eexists. split. eapply Smallstep.plus_one. eapply HTL.step_module; eauto. + apply assumption_32bit. econstructor. econstructor. econstructor. simplify. econstructor. econstructor. econstructor. simplify. eapply Verilog.erun_Vbinop with (EQ := ?[EQ1]). (* FIXME: These will be shelved and cause sadness. *) @@ -859,6 +861,7 @@ Section CORRECTNESS. eexists. split. eapply Smallstep.plus_one. eapply HTL.step_module; eauto. + apply assumption_32bit. econstructor. econstructor. econstructor. simplify. econstructor. econstructor. econstructor. simplify. eapply Verilog.erun_Vbinop with (EQ := ?[EQ3]). (* FIXME: These will be shelved and cause sadness. *) @@ -1037,6 +1040,7 @@ Section CORRECTNESS. eexists. split. eapply Smallstep.plus_one. eapply HTL.step_module; eauto. + apply assumption_32bit. econstructor. econstructor. econstructor. simplify. econstructor. econstructor. econstructor. econstructor. simplify. @@ -1201,6 +1205,7 @@ Section CORRECTNESS. eexists. split. eapply Smallstep.plus_one. eapply HTL.step_module; eauto. + apply assumption_32bit. econstructor. econstructor. econstructor. eapply Verilog.stmnt_runp_Vnonblock_arr. simplify. econstructor. @@ -1454,6 +1459,7 @@ Section CORRECTNESS. eexists. split. eapply Smallstep.plus_one. eapply HTL.step_module; eauto. + apply assumption_32bit. econstructor. econstructor. econstructor. eapply Verilog.stmnt_runp_Vnonblock_arr. simplify. econstructor. @@ -1672,6 +1678,7 @@ Section CORRECTNESS. eexists. split. eapply Smallstep.plus_one. eapply HTL.step_module; eauto. + apply assumption_32bit. econstructor. econstructor. econstructor. eapply Verilog.stmnt_runp_Vnonblock_arr. simplify. econstructor. econstructor. econstructor. econstructor. @@ -1856,6 +1863,7 @@ Section CORRECTNESS. - eexists. split. apply Smallstep.plus_one. eapply HTL.step_module; eauto. + apply assumption_32bit. eapply Verilog.stmnt_runp_Vnonblock_reg with (rhsval := if b then posToValue 32 ifso else posToValue 32 ifnot). @@ -1966,6 +1974,7 @@ Section CORRECTNESS. eapply Smallstep.plus_two. eapply HTL.step_module; eauto. + apply assumption_32bit. constructor. econstructor; simpl; trivial. econstructor; simpl; trivial. @@ -1998,6 +2007,7 @@ Section CORRECTNESS. - econstructor. split. eapply Smallstep.plus_two. eapply HTL.step_module; eauto. + apply assumption_32bit. constructor. econstructor; simpl; trivial. econstructor; simpl; trivial. diff --git a/src/verilog/Value.v b/src/verilog/Value.v index 52a87e3..e7b2362 100644 --- a/src/verilog/Value.v +++ b/src/verilog/Value.v @@ -345,6 +345,18 @@ Proof. simpl. rewrite <- Pos2Z.inj_pow_pos. trivial. Qed. +Lemma valueToPos_posToValue : + forall p, valueToPos (posToValueAuto p) = p. +Proof. + intros. unfold valueToPos, posToValueAuto. + rewrite uvalueToZ_ZToValue. auto. rewrite positive_nat_Z. + split. apply Zle_0_pos. + + assert (p < 2 ^ (Pos.size p))%positive by apply Pos.size_gt. + 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. -- cgit From 9aa32499597678e3b0e7ef0b8a85ca5beda44938 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Mon, 29 Jun 2020 12:21:38 +0100 Subject: Add missing file. --- src/common/ZExtra.v | 34 ++++++++++++++++++++++++++++++++++ 1 file changed, 34 insertions(+) create mode 100644 src/common/ZExtra.v diff --git a/src/common/ZExtra.v b/src/common/ZExtra.v new file mode 100644 index 0000000..a0dd717 --- /dev/null +++ b/src/common/ZExtra.v @@ -0,0 +1,34 @@ +Require Import ZArith. +Require Import Lia. + +Local Open Scope Z_scope. + +Module ZExtra. + + Lemma mod_0_bounds : + forall x y m, + 0 < m -> + x mod m = 0 -> + y mod m = 0 -> + x <> y -> + x + m > y -> + y + m <= x. + Proof. + intros x y m. + intros POS XMOD YMOD NEQ H. + destruct (Z_le_gt_dec (y + m) x); eauto. + + apply Znumtheory.Zmod_divide in YMOD; try lia. + apply Znumtheory.Zmod_divide in XMOD; try lia. + inversion XMOD as [x']; subst; clear XMOD. + inversion YMOD as [y']; subst; clear YMOD. + + assert (x' <> y') as NEQ' by lia; clear NEQ. + + rewrite <- Z.mul_succ_l in H. + rewrite <- Z.mul_succ_l in g. + apply Zmult_gt_reg_r in H; + apply Zmult_gt_reg_r in g; lia. + Qed. + +End ZExtra. -- cgit From 8398615fea7ab754854cb10e16e86de6415f1f2d Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Mon, 29 Jun 2020 16:55:08 +0100 Subject: Work on addition proof --- src/verilog/Value.v | 22 +++++++++++++++++++++- 1 file changed, 21 insertions(+), 1 deletion(-) diff --git a/src/verilog/Value.v b/src/verilog/Value.v index e7b2362..c380ca7 100644 --- a/src/verilog/Value.v +++ b/src/verilog/Value.v @@ -108,6 +108,12 @@ Definition boolToValue (sz : nat) (b : bool) : value := Definition unify_word (sz1 sz2 : nat) (w1 : word sz2): sz1 = sz2 -> word sz1. intros; subst; assumption. Defined. +Lemma unify_word_unfold : + forall sz w, + unify_word sz sz w eq_refl = w. +Proof. + intros. unfold unify_word. Admitted. + Definition value_eq_size: forall v1 v2 : value, { vsize v1 = vsize v2 } + { True }. Proof. @@ -382,7 +388,21 @@ Qed. Lemma boolToValue_ValueToBool : forall b, valueToBool (boolToValue 32 b) = b. -Proof. destruct b; unfold valueToBool, boolToValue; simpl; trivial. Qed. +Proof. destruct b; auto. Qed. + +Lemma intToValue_eq_size : + forall n1 n2, + vsize (intToValue n1) = vsize (intToValue n2). +Proof. auto. Qed. + +Local Open Scope Z. + +Lemma zadd_vplus : + forall z1 z2, + valueToZ (vplus (ZToValue 32 z1) (ZToValue 32 z2) eq_refl) = z1 + z2. +Proof. + intros. unfold valueToZ, ZToValue. simpl. + Admitted. (*Lemma ZToValue_valueToNat : forall x sz, -- cgit From 7e59d2723fb9c5b4631f5eac1e99ae8956871a7f Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Mon, 29 Jun 2020 16:59:31 +0100 Subject: Develop compiles again --- Makefile | 1 + src/translation/HTLgen.v | 10 +++++----- src/translation/HTLgenproof.v | 2 +- src/translation/Veriloggenproof.v | 6 ++++-- 4 files changed, 11 insertions(+), 8 deletions(-) diff --git a/Makefile b/Makefile index 3d31b2f..1303b13 100644 --- a/Makefile +++ b/Makefile @@ -67,3 +67,4 @@ clean:: Makefile.coq clean:: rm -f */*.v.d */*.glob */*.vo */*~ *~ + rm -f src/extraction/*.ml src/extraction/*.mli diff --git a/src/translation/HTLgen.v b/src/translation/HTLgen.v index 59fb70a..b32ed9d 100644 --- a/src/translation/HTLgen.v +++ b/src/translation/HTLgen.v @@ -294,24 +294,24 @@ Definition translate_eff_addressing (a: Op.addressing) (args: list reg) : mon ex | Op.Aindexed off, r1::nil => if (check_address_parameter_signed off) then ret (boplitz Vadd r1 off) - else error (Errors.msg "Veriloggen: translate_eff_addressing address misaligned") + else error (Errors.msg "Veriloggen: translate_eff_addressing (Aindexed): address misaligned") | Op.Ascaled scale offset, r1::nil => if (check_address_parameter_signed scale) && (check_address_parameter_signed offset) then ret (Vbinop Vadd (boplitz Vmul r1 scale) (Vlit (ZToValue 32 offset))) - else error (Errors.msg "Veriloggen: translate_eff_addressing address misaligned") + else error (Errors.msg "Veriloggen: translate_eff_addressing (Ascaled): address misaligned") | Op.Aindexed2 offset, r1::r2::nil => if (check_address_parameter_signed offset) then ret (Vbinop Vadd (Vvar r1) (boplitz Vadd r2 offset)) - else error (Errors.msg "Veriloggen: translate_eff_addressing address misaligned") + else error (Errors.msg "Veriloggen: translate_eff_addressing (Aindexed2): address misaligned") | Op.Aindexed2scaled scale offset, r1::r2::nil => (* Typical for dynamic array addressing *) if (check_address_parameter_signed scale) && (check_address_parameter_signed offset) then ret (Vbinop Vadd (boplitz Vadd r1 offset) (boplitz Vmul r2 scale)) - else error (Errors.msg "Veriloggen: translate_eff_addressing address misaligned") + else error (Errors.msg "Veriloggen: translate_eff_addressing (Aindexed2scaled): address misaligned") | Op.Ainstack a, nil => (* We need to be sure that the base address is aligned *) let a := Integers.Ptrofs.unsigned a in if (check_address_parameter_unsigned a) then ret (Vlit (ZToValue 32 a)) - else error (Errors.msg "Veriloggen: translate_eff_addressing address misaligned") + else error (Errors.msg "Veriloggen: translate_eff_addressing (Ainstack): address misaligned") | _, _ => error (Errors.msg "Veriloggen: translate_eff_addressing unsuported addressing") end. diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index 2f296f2..6dd0688 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -18,7 +18,7 @@ From compcert Require RTL Registers AST Integers. From compcert Require Import Globalenvs Memory. -From coqup Require Import Coquplib HTLgenspec HTLgen Value AssocMap Array IntegerExtra. +From coqup Require Import Coquplib HTLgenspec HTLgen Value AssocMap Array IntegerExtra ZExtra. From coqup Require HTL Verilog. Require Import Lia. diff --git a/src/translation/Veriloggenproof.v b/src/translation/Veriloggenproof.v index db96949..ca4ecab 100644 --- a/src/translation/Veriloggenproof.v +++ b/src/translation/Veriloggenproof.v @@ -69,12 +69,14 @@ Section CORRECTNESS. match_states S1 R1 -> exists R2, Smallstep.plus step tge R1 t R2 /\ match_states S2 R2. Proof. - induction 1; intros R1 MSTATE; inv MSTATE; econstructor; split. +(* induction 1; intros R1 MSTATE; inv MSTATE; econstructor; split. - apply Smallstep.plus_one. econstructor. eassumption. trivial. - * econstructor. econstructor. + * econstructor. econstructor.*) + Admitted. Theorem transf_program_correct: forward_simulation (HTL.semantics prog) (Verilog.semantics tprog). + Admitted. End CORRECTNESS. -- cgit From 1e0d5047d2272fdeb06391d1c5fa4e0472be2365 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Mon, 29 Jun 2020 21:29:32 +0100 Subject: Eliminate memory bounds assumption! --- src/common/IntegerExtra.v | 4 +- src/translation/HTLgenproof.v | 244 ++++++++++++++++++++++++++++++++++++------ 2 files changed, 211 insertions(+), 37 deletions(-) diff --git a/src/common/IntegerExtra.v b/src/common/IntegerExtra.v index 8df70d9..7d3156b 100644 --- a/src/common/IntegerExtra.v +++ b/src/common/IntegerExtra.v @@ -143,7 +143,7 @@ Module PtrofsExtra. Lemma divu_unsigned : forall x y, 0 < Ptrofs.unsigned y -> - Ptrofs.unsigned x < Ptrofs.max_unsigned -> + Ptrofs.unsigned x <= Ptrofs.max_unsigned -> Ptrofs.unsigned (Ptrofs.divu x y) = Ptrofs.unsigned x / Ptrofs.unsigned y. Proof. intros. @@ -154,7 +154,7 @@ Module PtrofsExtra. apply Ptrofs.unsigned_range. apply Z.div_le_upper_bound; auto. eapply Z.le_trans. - apply Z.lt_le_incl. exact H0. + exact H0. rewrite Z.mul_comm. apply Z.le_mul_diag_r; simplify; lia. Qed. diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index 2f296f2..f5a55af 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -18,7 +18,7 @@ From compcert Require RTL Registers AST Integers. From compcert Require Import Globalenvs Memory. -From coqup Require Import Coquplib HTLgenspec HTLgen Value AssocMap Array IntegerExtra. +From coqup Require Import Coquplib HTLgenspec HTLgen Value AssocMap Array IntegerExtra ZExtra. From coqup Require HTL Verilog. Require Import Lia. @@ -76,6 +76,13 @@ Definition arr_stack_based_pointers (spb : Values.block) (m : mem) (stack_length (Values.Val.offset_ptr sp (Integers.Ptrofs.repr (4 * ptr))))) spb. +Definition stack_bounds (sp : Values.val) (hi : Z) (m : mem) : Prop := + forall ptr v, + hi <= ptr <= Integers.Ptrofs.max_unsigned -> + Z.modulo ptr 4 = 0 -> + Mem.loadv AST.Mint32 m (Values.Val.offset_ptr sp (Integers.Ptrofs.repr ptr )) = None /\ + Mem.storev AST.Mint32 m (Values.Val.offset_ptr sp (Integers.Ptrofs.repr ptr )) v = None. + Inductive match_stacks (mem : mem) : list RTL.stackframe -> list HTL.stackframe -> Prop := | match_stacks_nil : match_stacks mem nil nil @@ -87,7 +94,8 @@ Inductive match_stacks (mem : mem) : list RTL.stackframe -> list HTL.stackframe (MARR : match_arrs m f sp mem asa) (SP : sp = Values.Vptr sp' (Integers.Ptrofs.repr 0)) (RSBP: reg_stack_based_pointers sp' rs) - (ASBP: arr_stack_based_pointers sp' mem (f.(RTL.fn_stacksize)) sp), + (ASBP: arr_stack_based_pointers sp' mem (f.(RTL.fn_stacksize)) sp) + (BOUNDS : stack_bounds sp (f.(RTL.fn_stacksize)) mem), match_stacks mem (RTL.Stackframe r f sp pc rs :: cs) (HTL.Stackframe r m pc asr asa :: lr). @@ -99,8 +107,9 @@ Inductive match_states : RTL.state -> HTL.state -> Prop := (MS : match_stacks mem sf res) (MARR : match_arrs m f sp mem asa) (SP : sp = Values.Vptr sp' (Integers.Ptrofs.repr 0)) - (RSBP: reg_stack_based_pointers sp' rs) - (ASBP: arr_stack_based_pointers sp' mem (f.(RTL.fn_stacksize)) sp), + (RSBP : reg_stack_based_pointers sp' rs) + (ASBP : arr_stack_based_pointers sp' mem (f.(RTL.fn_stacksize)) sp) + (BOUNDS : stack_bounds sp (f.(RTL.fn_stacksize)) mem), match_states (RTL.State sf f sp st rs mem) (HTL.State res m st asr asa) | match_returnstate : @@ -601,9 +610,6 @@ Section CORRECTNESS. remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0)) (Integers.Ptrofs.of_int (Integers.Int.repr z))) as OFFSET. - (** Read bounds assumption *) - assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH by admit. - (** Modular preservation proof *) assert (Integers.Ptrofs.unsigned OFFSET mod 4 = 0) as MOD_PRESERVE. { rewrite HeqOFFSET. @@ -615,6 +621,18 @@ Section CORRECTNESS. rewrite Integers.Int.signed_repr; simplify; try split; try assumption. } + (** Read bounds proof *) + assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH. + { destruct (Integers.Ptrofs.unsigned OFFSET destruct x eqn:EQ end; try reflexivity. + assert (Mem.valid_access m AST.Mint32 sp' + (Integers.Ptrofs.unsigned + (Integers.Ptrofs.add (Integers.Ptrofs.repr 0) + (Integers.Ptrofs.repr ptr))) Writable). + { pose proof H1. eapply Mem.store_valid_access_2 in H0. + exact H0. eapply Mem.store_valid_access_3. eassumption. } + pose proof (Mem.valid_access_store m AST.Mint32 sp' + (Integers.Ptrofs.unsigned + (Integers.Ptrofs.add (Integers.Ptrofs.repr 0) + (Integers.Ptrofs.repr ptr))) v). + apply X in H0. invert H0. congruence. + + (** Preamble *) invert MARR. simplify. @@ -1435,9 +1505,6 @@ Section CORRECTNESS. (Integers.Int.add (Integers.Int.mul (valueToInt asr # r1) (Integers.Int.repr z)) (Integers.Int.repr z0)))) as OFFSET. - (** Read bounds assumption *) - assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH by admit. - (** Modular preservation proof *) assert (Integers.Ptrofs.unsigned OFFSET mod 4 = 0) as MOD_PRESERVE. { rewrite HeqOFFSET. @@ -1455,6 +1522,18 @@ Section CORRECTNESS. rewrite Integers.Int.signed_repr; simplify; try split; try assumption. } + (** Write bounds proof *) + assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as WRITE_BOUND_HIGH. + { destruct (Integers.Ptrofs.unsigned OFFSET destruct x eqn:EQ end; try reflexivity. + assert (Mem.valid_access m AST.Mint32 sp' + (Integers.Ptrofs.unsigned + (Integers.Ptrofs.add (Integers.Ptrofs.repr 0) + (Integers.Ptrofs.repr ptr))) Writable). + { pose proof H1. eapply Mem.store_valid_access_2 in H0. + exact H0. eapply Mem.store_valid_access_3. eassumption. } + pose proof (Mem.valid_access_store m AST.Mint32 sp' + (Integers.Ptrofs.unsigned + (Integers.Ptrofs.add (Integers.Ptrofs.repr 0) + (Integers.Ptrofs.repr ptr))) v). + apply X in H0. invert H0. congruence. + + invert MARR. simplify. unfold Op.eval_addressing in H0. @@ -1668,12 +1774,20 @@ Section CORRECTNESS. remember i0 as OFFSET. - (** Read bounds assumption *) - assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH by admit. - (** Modular preservation proof *) rename H0 into MOD_PRESERVE. + (** Write bounds proof *) + assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as WRITE_BOUND_HIGH. + { destruct (Integers.Ptrofs.unsigned OFFSET destruct x eqn:EQ end; try reflexivity. + assert (Mem.valid_access m AST.Mint32 sp' + (Integers.Ptrofs.unsigned + (Integers.Ptrofs.add (Integers.Ptrofs.repr 0) + (Integers.Ptrofs.repr ptr))) Writable). + { pose proof H1. eapply Mem.store_valid_access_2 in H0. + exact H0. eapply Mem.store_valid_access_3. eassumption. } + pose proof (Mem.valid_access_store m AST.Mint32 sp' + (Integers.Ptrofs.unsigned + (Integers.Ptrofs.add (Integers.Ptrofs.repr 0) + (Integers.Ptrofs.repr ptr))) v). + apply X in H0. invert H0. congruence. + - eexists. split. apply Smallstep.plus_one. eapply HTL.step_module; eauto. apply assumption_32bit. @@ -2090,6 +2231,49 @@ Section CORRECTNESS. repeat constructor. constructor. + Transparent Mem.alloc. (* TODO: Since there are opaque there's probably a lemma. *) + Transparent Mem.load. + Transparent Mem.store. + unfold stack_bounds. + split. + + unfold Mem.alloc in H. + invert H. + simplify. + unfold Mem.load. + intros. + match goal with | |- context[if ?x then _ else _] => destruct x end; try congruence. + invert v0. unfold Mem.range_perm in H3. + unfold Mem.perm in H3. simplify. + unfold Mem.perm_order' in H3. + rewrite Integers.Ptrofs.add_zero_l in H3. + rewrite Integers.Ptrofs.unsigned_repr in H3; simplify; try lia. + exploit (H3 ptr). lia. intros. + rewrite Maps.PMap.gss in H8. + match goal with | H8 : context[if ?x then _ else _] |- _ => destruct x eqn:EQ end; try contradiction. + simplify. + apply proj_sumbool_true in H10. lia. + + unfold Mem.alloc in H. + invert H. + simplify. + unfold Mem.store. + intros. + match goal with | |- context[if ?x then _ else _] => destruct x end; try congruence. + invert v0. unfold Mem.range_perm in H3. + unfold Mem.perm in H3. simplify. + unfold Mem.perm_order' in H3. + rewrite Integers.Ptrofs.add_zero_l in H3. + rewrite Integers.Ptrofs.unsigned_repr in H3; simplify; try lia. + exploit (H3 ptr). lia. intros. + rewrite Maps.PMap.gss in H8. + match goal with | H8 : context[if ?x then _ else _] |- _ => destruct x eqn:EQ end; try contradiction. + simplify. + apply proj_sumbool_true in H10. lia. + Opaque Mem.alloc. + Opaque Mem.load. + Opaque Mem.store. + - invert MSTATE. invert MS. econstructor. split. apply Smallstep.plus_one. @@ -2104,19 +2288,9 @@ Section CORRECTNESS. rewrite AssocMap.gss. trivial. apply st_greater_than_res. admit. - - Unshelve. - exact (AssocMap.empty value). - exact (AssocMap.empty value). - (* exact (ZToValue 32 0). *) - (* exact (AssocMap.empty value). *) - (* exact (AssocMap.empty value). *) - (* exact (AssocMap.empty value). *) - (* exact (AssocMap.empty value). *) Admitted. Hint Resolve transl_step_correct : htlproof. - Lemma option_inv : forall A x y, @Some A x = Some y -> x = y. -- cgit From 8a260a8b81617e192fc929e0189fd3df1327f80a Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Tue, 30 Jun 2020 00:04:39 +0100 Subject: Add equivalence between int add value add --- src/verilog/Value.v | 64 ++++++++++++++++++++++++++++++++++++++++++++--------- 1 file changed, 54 insertions(+), 10 deletions(-) diff --git a/src/verilog/Value.v b/src/verilog/Value.v index c380ca7..3c5ed04 100644 --- a/src/verilog/Value.v +++ b/src/verilog/Value.v @@ -111,8 +111,7 @@ intros; subst; assumption. Defined. Lemma unify_word_unfold : forall sz w, unify_word sz sz w eq_refl = w. -Proof. - intros. unfold unify_word. Admitted. +Proof. auto. Qed. Definition value_eq_size: forall v1 v2 : value, { vsize v1 = vsize v2 } + { True }. @@ -330,6 +329,12 @@ Proof. auto using uwordToZ_ZToWord. Qed. +Lemma uvalueToZ_ZToValue_full : + forall sz : nat, + (0 < sz)%nat -> + forall z : Z, uvalueToZ (ZToValue sz z) = (z mod 2 ^ Z.of_nat sz)%Z. +Proof. unfold uvalueToZ, ZToValue. simpl. auto using uwordToZ_ZToWord_full. Qed. + Lemma ZToValue_uvalueToZ : forall v, ZToValue (vsize v) (uvalueToZ v) = v. @@ -390,19 +395,58 @@ Lemma boolToValue_ValueToBool : valueToBool (boolToValue 32 b) = b. Proof. destruct b; auto. Qed. -Lemma intToValue_eq_size : - forall n1 n2, - vsize (intToValue n1) = vsize (intToValue n2). -Proof. auto. Qed. - Local Open Scope Z. +Ltac eq_op H := + intros; unfold uvalueToZ, ZToValue; simpl; rewrite unify_word_unfold; + rewrite <- H; rewrite uwordToZ_ZToWord; auto. + Lemma zadd_vplus : forall z1 z2, - valueToZ (vplus (ZToValue 32 z1) (ZToValue 32 z2) eq_refl) = z1 + z2. + 0 <= z1 + z2 < 2 ^ 32 -> + uvalueToZ (vplus (ZToValue 32 z1) (ZToValue 32 z2) eq_refl) = z1 + z2. +Proof. eq_op ZToWord_plus. Qed. + +Lemma zadd_vplus2 : + forall z1 z2, + vplus (ZToValue 32 z1) (ZToValue 32 z2) eq_refl = ZToValue 32 (z1 + z2). Proof. - intros. unfold valueToZ, ZToValue. simpl. - Admitted. + intros. unfold vplus, ZToValue, map_word2. rewrite unify_word_unfold. simpl. + rewrite ZToWord_plus; auto. +Qed. + +Lemma intadd_vplus : + forall i1 i2, + valueToInt (vplus (intToValue i1) (intToValue i2) eq_refl) = Int.add i1 i2. +Proof. + intros. unfold Int.add, valueToInt, intToValue. rewrite zadd_vplus2. + rewrite uvalueToZ_ZToValue_full. rewrite <- Int.unsigned_repr_eq. + rewrite Int.repr_unsigned. auto. omega. +Qed. + +Lemma zsub_vminus : + forall z1 z2, + 0 <= z1 - z2 < 2 ^ 32 -> + uvalueToZ (vminus (ZToValue 32 z1) (ZToValue 32 z2) eq_refl) = z1 - z2. +Proof. eq_op ZToWord_minus. Qed. + +Lemma zmul_vmul : + forall z1 z2, + 0 <= z1 * z2 < 2 ^ 32 -> + uvalueToZ (vmul (ZToValue 32 z1) (ZToValue 32 z2) eq_refl) = z1 * z2. +Proof. eq_op ZToWord_mult. Qed. + +Local Open Scope N. +Lemma zdiv_vdiv : + forall n1 n2, + n1 < 2 ^ 32 -> + n2 < 2 ^ 32 -> + n1 / n2 < 2 ^ 32 -> + valueToN (vdiv (NToValue 32 n1) (NToValue 32 n2) eq_refl) = n1 / n2. +Proof. + intros; unfold valueToN, NToValue; simpl; rewrite unify_word_unfold. unfold wdiv. + unfold wordBin. repeat (rewrite wordToN_NToWord_2); auto. +Qed. (*Lemma ZToValue_valueToNat : forall x sz, -- cgit From 9a640d4f50b93a9d5f9b48e10b0c4cbfa7f0680f Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Tue, 30 Jun 2020 00:29:41 +0100 Subject: Make the proofs more concise --- src/verilog/Value.v | 38 +++++++++++++++++++++----------------- 1 file changed, 21 insertions(+), 17 deletions(-) diff --git a/src/verilog/Value.v b/src/verilog/Value.v index 3c5ed04..8ba5138 100644 --- a/src/verilog/Value.v +++ b/src/verilog/Value.v @@ -397,15 +397,15 @@ Proof. destruct b; auto. Qed. Local Open Scope Z. -Ltac eq_op H := +Ltac word_op_value H := intros; unfold uvalueToZ, ZToValue; simpl; rewrite unify_word_unfold; - rewrite <- H; rewrite uwordToZ_ZToWord; auto. + rewrite <- H; rewrite uwordToZ_ZToWord_full; auto; omega. Lemma zadd_vplus : - forall z1 z2, - 0 <= z1 + z2 < 2 ^ 32 -> - uvalueToZ (vplus (ZToValue 32 z1) (ZToValue 32 z2) eq_refl) = z1 + z2. -Proof. eq_op ZToWord_plus. Qed. + forall sz z1 z2, + (sz > 0)%nat -> + uvalueToZ (vplus (ZToValue sz z1) (ZToValue sz z2) eq_refl) = (z1 + z2) mod 2 ^ Z.of_nat sz. +Proof. word_op_value ZToWord_plus. Qed. Lemma zadd_vplus2 : forall z1 z2, @@ -415,26 +415,30 @@ Proof. rewrite ZToWord_plus; auto. Qed. +Lemma wordsize_32 : + Int.wordsize = 32%nat. +Proof. auto. Qed. + Lemma intadd_vplus : forall i1 i2, valueToInt (vplus (intToValue i1) (intToValue i2) eq_refl) = Int.add i1 i2. Proof. - intros. unfold Int.add, valueToInt, intToValue. rewrite zadd_vplus2. - rewrite uvalueToZ_ZToValue_full. rewrite <- Int.unsigned_repr_eq. - rewrite Int.repr_unsigned. auto. omega. + intros. unfold Int.add, valueToInt, intToValue. rewrite zadd_vplus. + rewrite <- Int.unsigned_repr_eq. + rewrite Int.repr_unsigned. auto. rewrite wordsize_32. omega. Qed. Lemma zsub_vminus : - forall z1 z2, - 0 <= z1 - z2 < 2 ^ 32 -> - uvalueToZ (vminus (ZToValue 32 z1) (ZToValue 32 z2) eq_refl) = z1 - z2. -Proof. eq_op ZToWord_minus. Qed. + forall sz z1 z2, + (sz > 0)%nat -> + uvalueToZ (vminus (ZToValue sz z1) (ZToValue sz z2) eq_refl) = (z1 - z2) mod 2 ^ Z.of_nat sz. +Proof. word_op_value ZToWord_minus. Qed. Lemma zmul_vmul : - forall z1 z2, - 0 <= z1 * z2 < 2 ^ 32 -> - uvalueToZ (vmul (ZToValue 32 z1) (ZToValue 32 z2) eq_refl) = z1 * z2. -Proof. eq_op ZToWord_mult. Qed. + forall sz z1 z2, + (sz > 0)%nat -> + uvalueToZ (vmul (ZToValue sz z1) (ZToValue sz z2) eq_refl) = (z1 * z2) mod 2 ^ Z.of_nat sz. +Proof. word_op_value ZToWord_mult. Qed. Local Open Scope N. Lemma zdiv_vdiv : -- cgit From f26f3887d0b0ac286c317a5425a3a4781871cfc2 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Tue, 30 Jun 2020 12:34:35 +0100 Subject: Add command line flags for initial block --- driver/CoqupDriver.ml | 8 +++----- src/CoqupClflags.ml | 5 +++++ src/verilog/PrintVerilog.ml | 10 ++++++++++ test/loop.c | 6 ++++-- 4 files changed, 22 insertions(+), 7 deletions(-) create mode 100644 src/CoqupClflags.ml diff --git a/driver/CoqupDriver.ml b/driver/CoqupDriver.ml index 2932a50..afbe4d0 100644 --- a/driver/CoqupDriver.ml +++ b/driver/CoqupDriver.ml @@ -36,11 +36,7 @@ open Coqup.Frontend open Coqup.Assembler open Coqup.Linker open Coqup.Diagnostics - -(* Coqup flags *) -let option_simulate = ref false -let option_hls = ref true -let option_debug_hls = ref false +open Coqup.CoqupClflags (* Name used for version string etc. *) let tool_name = "C verified high-level synthesis" @@ -215,6 +211,7 @@ Processing options: --no-hls Do not use HLS and generate standard flow. --simulate Simulate the result with the Verilog semantics. --debug-hls Add debug logic to the Verilog. + --initialise-stack initialise the stack to all 0s. |} ^ prepro_help ^ language_support_help ^ @@ -316,6 +313,7 @@ let cmdline_actions = [Exact "--no-hls", Unset option_hls; Exact "--simulate", Set option_simulate; Exact "--debug-hls", Set option_debug_hls; + Exact "--initialise-stack", Set option_initial; ] (* Getting version info *) @ version_options tool_name @ diff --git a/src/CoqupClflags.ml b/src/CoqupClflags.ml new file mode 100644 index 0000000..83dd31d --- /dev/null +++ b/src/CoqupClflags.ml @@ -0,0 +1,5 @@ +(* Coqup flags *) +let option_simulate = ref false +let option_hls = ref true +let option_debug_hls = ref false +let option_initial = ref false diff --git a/src/verilog/PrintVerilog.ml b/src/verilog/PrintVerilog.ml index 5dc0386..6d10887 100644 --- a/src/verilog/PrintVerilog.ml +++ b/src/verilog/PrintVerilog.ml @@ -22,9 +22,12 @@ open Datatypes open Camlcoq open AST +open Clflags open Printf +open CoqupClflags + let concat = String.concat "" let indent i = String.make (2 * i) ' ' @@ -187,6 +190,12 @@ let debug_always i clk state = concat [ indent i; "end\n" ] +let print_initial i n stk = concat [ + indent i; "integer i;\n"; + indent i; "initial for(i = 0; i < "; sprintf "%d" n; "; i++)\n"; + indent (i+1); register stk; "[i] = 0;\n" + ] + let pprint_module debug i n m = if (extern_atom n) = "main" then let inputs = m.mod_start :: m.mod_reset :: m.mod_clk :: m.mod_args in @@ -194,6 +203,7 @@ let pprint_module debug i n m = concat [ indent i; "module "; (extern_atom n); "("; concat (intersperse ", " (List.map register (inputs @ outputs))); ");\n"; fold_map (pprint_module_item (i+1)) m.mod_body; + if !option_initial then print_initial i (Nat.to_int m.mod_stk_len) m.mod_stk else ""; if debug then debug_always i m.mod_clk m.mod_st else ""; indent i; "endmodule\n\n" ] diff --git a/test/loop.c b/test/loop.c index b459e3a..52e4fe9 100644 --- a/test/loop.c +++ b/test/loop.c @@ -1,10 +1,12 @@ int main() { int max = 5; int acc = 0; + int b = 1; + int c = 2; - for (int i = 0; i < max; i++) { + for (int i = 0; i < max; i = i + b) { acc += i; } - return acc + 2; + return acc + c; } -- cgit From 7d19237389edebd7cc897494f7db2a4c8dcc97b4 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Tue, 30 Jun 2020 16:38:33 +0100 Subject: Fix stack frame issue. We never cons a stack frame since we don't support calls (aside from the initial call which doesn't push a stack frame); removing the cons constructor solves the issue regarding memory separation. This means we now _can't_ support calls even if we wanted to, but due to the way we implement memory, we would need quite a lot of extra work to support this. --- src/translation/HTLgenproof.v | 71 +++++++++++++++---------------------------- 1 file changed, 25 insertions(+), 46 deletions(-) diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index f5a55af..3aff5c9 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -83,28 +83,28 @@ Definition stack_bounds (sp : Values.val) (hi : Z) (m : mem) : Prop := Mem.loadv AST.Mint32 m (Values.Val.offset_ptr sp (Integers.Ptrofs.repr ptr )) = None /\ Mem.storev AST.Mint32 m (Values.Val.offset_ptr sp (Integers.Ptrofs.repr ptr )) v = None. -Inductive match_stacks (mem : mem) : list RTL.stackframe -> list HTL.stackframe -> Prop := -| match_stacks_nil : - match_stacks mem nil nil -| match_stacks_cons : - forall cs lr r f sp sp' pc rs m asr asa - (TF : tr_module f m) - (ST: match_stacks mem cs lr) - (MA: match_assocmaps f rs asr) - (MARR : match_arrs m f sp mem asa) - (SP : sp = Values.Vptr sp' (Integers.Ptrofs.repr 0)) - (RSBP: reg_stack_based_pointers sp' rs) - (ASBP: arr_stack_based_pointers sp' mem (f.(RTL.fn_stacksize)) sp) - (BOUNDS : stack_bounds sp (f.(RTL.fn_stacksize)) mem), - match_stacks mem (RTL.Stackframe r f sp pc rs :: cs) - (HTL.Stackframe r m pc asr asa :: lr). +Inductive match_frames : list RTL.stackframe -> list HTL.stackframe -> Prop := +| match_frames_nil : + match_frames nil nil. +(* | match_frames_cons : *) +(* forall cs lr r f sp sp' pc rs m asr asa *) +(* (TF : tr_module f m) *) +(* (ST: match_frames mem cs lr) *) +(* (MA: match_assocmaps f rs asr) *) +(* (MARR : match_arrs m f sp mem asa) *) +(* (SP : sp = Values.Vptr sp' (Integers.Ptrofs.repr 0)) *) +(* (RSBP: reg_stack_based_pointers sp' rs) *) +(* (ASBP: arr_stack_based_pointers sp' mem (f.(RTL.fn_stacksize)) sp) *) +(* (BOUNDS : stack_bounds sp (f.(RTL.fn_stacksize)) mem), *) +(* match_frames mem (RTL.Stackframe r f sp pc rs :: cs) *) +(* (HTL.Stackframe r m pc asr asa :: lr). *) Inductive match_states : RTL.state -> HTL.state -> Prop := | match_state : forall asa asr sf f sp sp' rs mem m st res (MASSOC : match_assocmaps f rs asr) (TF : tr_module f m) (WF : state_st_wf m (HTL.State res m st asr asa)) - (MS : match_stacks mem sf res) + (MF : match_frames sf res) (MARR : match_arrs m f sp mem asa) (SP : sp = Values.Vptr sp' (Integers.Ptrofs.repr 0)) (RSBP : reg_stack_based_pointers sp' rs) @@ -115,7 +115,7 @@ Inductive match_states : RTL.state -> HTL.state -> Prop := | match_returnstate : forall v v' stack mem res - (MS : match_stacks mem stack res), + (MF : match_frames stack res), val_value_lessdef v v' -> match_states (RTL.Returnstate stack v mem) (HTL.Returnstate res v') | match_initial_call : @@ -1283,9 +1283,6 @@ Section CORRECTNESS. unfold_merge. apply AssocMap.gss. - (** Match stacks *) - admit. - (** Equality proof *) assert (Z.to_nat (Integers.Ptrofs.unsigned @@ -1575,9 +1572,6 @@ Section CORRECTNESS. unfold_merge. apply AssocMap.gss. - (** Match stacks *) - admit. - (** Equality proof *) assert (Z.to_nat (Integers.Ptrofs.unsigned @@ -1820,9 +1814,6 @@ Section CORRECTNESS. unfold_merge. apply AssocMap.gss. - (** Match stacks *) - admit. - (** Equality proof *) assert (Z.to_nat (Integers.Ptrofs.unsigned @@ -2140,9 +2131,8 @@ Section CORRECTNESS. apply finish_not_return. apply AssocMap.gss. rewrite Events.E0_left. reflexivity. - constructor. - admit. + constructor; auto. constructor. - econstructor. split. @@ -2174,8 +2164,7 @@ Section CORRECTNESS. apply AssocMap.gss. rewrite Events.E0_left. trivial. - constructor. - admit. + constructor; auto. simpl. inversion MASSOC. subst. unfold find_assocmap, AssocMapExt.get_default. rewrite AssocMap.gso. @@ -2192,7 +2181,9 @@ Section CORRECTNESS. apply greater_than_max_func. apply init_reg_assoc_empty. unfold state_st_wf. - intros. inv H3. apply AssocMap.gss. constructor. + intros. inv H3. apply AssocMap.gss. + + constructor. econstructor. simplify. repeat split. unfold HTL.empty_stack. @@ -2274,20 +2265,8 @@ Section CORRECTNESS. Opaque Mem.load. Opaque Mem.store. - - invert MSTATE. invert MS. - econstructor. - split. apply Smallstep.plus_one. - constructor. - - constructor; auto. - econstructor; auto. - apply regs_lessdef_add_match; auto. - apply regs_lessdef_add_greater. apply greater_than_max_func. auto. - - unfold state_st_wf. intros. inv H. rewrite AssocMap.gso. - rewrite AssocMap.gss. trivial. apply st_greater_than_res. - - admit. + - inversion MSTATE. + inversion MF. Admitted. Hint Resolve transl_step_correct : htlproof. @@ -2361,7 +2340,7 @@ Section CORRECTNESS. Smallstep.final_state (RTL.semantics prog) s1 r -> Smallstep.final_state (HTL.semantics tprog) s2 r. Proof. - intros. inv H0. inv H. inv H4. inv MS. constructor. trivial. + intros. inv H0. inv H. inv H4. invert MF. constructor. reflexivity. Qed. Hint Resolve transl_final_states : htlproof. -- cgit From ffc978ec677f2f37ab8d8d1bf865cddadf087b81 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Tue, 30 Jun 2020 17:38:37 +0100 Subject: Factor out lemmas in main induction proof. --- src/translation/HTLgenproof.v | 3482 +++++++++++++++++++++-------------------- 1 file changed, 1813 insertions(+), 1669 deletions(-) diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index 3aff5c9..252119a 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -421,48 +421,593 @@ Section CORRECTNESS. exists asr' asa', HTL.step tge (HTL.State res m st asr asa) Events.E0 (HTL.State res m st asr' asa'). + Lemma transl_inop_correct: + forall (s : list RTL.stackframe) (f : RTL.function) (sp : Values.val) (pc : positive) + (rs : RTL.regset) (m : mem) (pc' : RTL.node), + (RTL.fn_code f) ! pc = Some (RTL.Inop pc') -> + forall R1 : HTL.state, + match_states (RTL.State s f sp pc rs m) R1 -> + exists R2 : HTL.state, + Smallstep.plus HTL.step tge R1 Events.E0 R2 /\ match_states (RTL.State s f sp pc' rs m) R2. + Proof. + intros s f sp pc rs m pc' H R1 MSTATE. + inv_state. + + unfold match_prog in TRANSL. + econstructor. + split. + apply Smallstep.plus_one. + eapply HTL.step_module; eauto. + apply assumption_32bit. + (* processing of state *) + econstructor. + simplify. + econstructor. + econstructor. + econstructor. + simplify. + + unfold Verilog.merge_regs. + unfold_merge. apply AssocMap.gss. + + (* prove match_state *) + rewrite assumption_32bit. + econstructor; simplify; eauto. + + unfold Verilog.merge_regs. + unfold_merge. simpl. apply regs_lessdef_add_greater. apply greater_than_max_func. + assumption. + unfold Verilog.merge_regs. + unfold state_st_wf. inversion 1. subst. unfold_merge. apply AssocMap.gss. + + (* prove match_arrs *) + invert MARR. simplify. + unfold HTL.empty_stack. simplify. unfold Verilog.merge_arrs. + econstructor. + simplify. repeat split. + + rewrite AssocMap.gcombine. + 2: { reflexivity. } + rewrite AssocMap.gss. + unfold Verilog.merge_arr. + setoid_rewrite H5. + reflexivity. + + rewrite combine_length. + unfold arr_repeat. simplify. + rewrite list_repeat_len. + reflexivity. + + unfold arr_repeat. simplify. + rewrite list_repeat_len; auto. + intros. + erewrite array_get_error_equal. + eauto. apply combine_none. - Theorem transl_step_correct: - forall (S1 : RTL.state) t S2, - RTL.step ge S1 t S2 -> - forall (R1 : HTL.state), - match_states S1 R1 -> - exists R2, Smallstep.plus HTL.step tge R1 t R2 /\ match_states S2 R2. + assumption. + + Unshelve. + constructor. + Qed. + Hint Resolve transl_inop_correct : htlproof. + + Lemma transl_iop_correct: + forall (s : list RTL.stackframe) (f : RTL.function) (sp : Values.val) (pc : positive) + (rs : Registers.Regmap.t Values.val) (m : mem) (op : Op.operation) (args : list Registers.reg) + (res0 : Registers.reg) (pc' : RTL.node) (v : Values.val), + (RTL.fn_code f) ! pc = Some (RTL.Iop op args res0 pc') -> + Op.eval_operation ge sp op (map (fun r : positive => Registers.Regmap.get r rs) args) m = Some v -> + forall R1 : HTL.state, + match_states (RTL.State s f sp pc rs m) R1 -> + exists R2 : HTL.state, + Smallstep.plus HTL.step tge R1 Events.E0 R2 /\ + match_states (RTL.State s f sp pc' (Registers.Regmap.set res0 v rs) m) R2. Proof. - induction 1; intros R1 MSTATE; try inv_state. - - (* Inop *) - unfold match_prog in TRANSL. - econstructor. - split. - apply Smallstep.plus_one. + intros s f sp pc rs m op args res0 pc' v H H0. + + (* Iop *) + (* destruct v eqn:?; *) + (* try ( *) + (* 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); *) + (* try (unfold Op.eval_addressing32 in H6; repeat (unfold_match H6); inversion H6; *) + (* unfold_func H3); *) + + (* inversion Heql; inversion MASSOC; subst; *) + (* assert (HPle : Ple r (RTL.max_reg_function f)) *) + (* by (eapply RTL.max_reg_function_use; eauto; simpl; auto); *) + (* apply H1 in HPle; inversion HPle; *) + (* rewrite H2 in *; discriminate *) + (* ). *) + + (* + 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. *) + + (* (* match_states *) *) + (* assert (pc' = valueToPos (posToValue 32 pc')). auto using assumption_32bit. *) + (* rewrite <- H1. *) + (* constructor. *) + (* unfold_merge. *) + (* 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. *) + (* assumption. *) + Admitted. + Hint Resolve transl_iop_correct : htlproof. + + Ltac tac := + repeat match goal with + | [ _ : error _ _ = OK _ _ _ |- _ ] => discriminate + | [ _ : context[if (?x && ?y) then _ else _] |- _ ] => + let EQ1 := fresh "EQ" in + let EQ2 := fresh "EQ" in + destruct x eqn:EQ1; destruct y eqn:EQ2; simpl in * + | [ _ : context[if ?x then _ else _] |- _ ] => + let EQ := fresh "EQ" in + destruct x eqn:EQ; simpl in * + | [ H : ret _ _ = _ |- _ ] => invert H + | [ _ : context[match ?x with | _ => _ end] |- _ ] => destruct x + end. + + Lemma transl_iload_correct: + forall (s : list RTL.stackframe) (f : RTL.function) (sp : Values.val) (pc : positive) + (rs : Registers.Regmap.t Values.val) (m : mem) (chunk : AST.memory_chunk) + (addr : Op.addressing) (args : list Registers.reg) (dst : Registers.reg) + (pc' : RTL.node) (a v : Values.val), + (RTL.fn_code f) ! pc = Some (RTL.Iload chunk addr args dst pc') -> + Op.eval_addressing ge sp addr (map (fun r : positive => Registers.Regmap.get r rs) args) = Some a -> + Mem.loadv chunk m a = Some v -> + forall R1 : HTL.state, + match_states (RTL.State s f sp pc rs m) R1 -> + exists R2 : HTL.state, + Smallstep.plus HTL.step tge R1 Events.E0 R2 /\ + match_states (RTL.State s f sp pc' (Registers.Regmap.set dst v rs) m) R2. + Proof. + intros s f sp pc rs m chunk addr args dst pc' a v H H0 H1 R1 MSTATE. + inv_state. + + destruct c, chunk, addr, args; simplify; tac; simplify. + + + (** Preamble *) + invert MARR. simplify. + + unfold Op.eval_addressing in H0. + destruct (Archi.ptr64) eqn:ARCHI; simplify. + + unfold reg_stack_based_pointers in RSBP. + pose proof (RSBP r0) as RSBPr0. + + destruct (Registers.Regmap.get r0 rs) eqn:EQr0; simplify. + + rewrite ARCHI in H1. simplify. + subst. + + pose proof MASSOC as MASSOC'. + invert MASSOC'. + pose proof (H0 r0). + assert (HPler0 : Ple r0 (RTL.max_reg_function f)) + by (eapply RTL.max_reg_function_use; eauto; simplify; eauto). + apply H6 in HPler0. + invert HPler0; try congruence. + rewrite EQr0 in H8. + invert H8. + clear H0. clear H6. + + unfold check_address_parameter_signed in *; + unfold check_address_parameter_unsigned in *; simplify. + + remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0)) + (Integers.Ptrofs.of_int (Integers.Int.repr z))) as OFFSET. + + (** Modular preservation proof *) + assert (Integers.Ptrofs.unsigned OFFSET mod 4 = 0) as MOD_PRESERVE. + { rewrite HeqOFFSET. + apply PtrofsExtra.add_mod; simplify; try lia. + exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *) + rewrite Integers.Ptrofs.signed_repr; try assumption. + admit. (* FIXME: Register bounds. *) + apply PtrofsExtra.of_int_mod. + rewrite Integers.Int.signed_repr; simplify; try split; try assumption. + } + + (** Read bounds proof *) + assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH. + { destruct (Integers.Ptrofs.unsigned OFFSET discriminate - | [ _ : context[if (?x && ?y) then _ else _] |- _ ] => - let EQ1 := fresh "EQ" in - let EQ2 := fresh "EQ" in - destruct x eqn:EQ1; destruct y eqn:EQ2; simpl in * - | [ _ : context[if ?x then _ else _] |- _ ] => - let EQ := fresh "EQ" in - destruct x eqn:EQ; simpl in * - | [ H : ret _ _ = _ |- _ ] => invert H - | [ _ : context[match ?x with | _ => _ end] |- _ ] => destruct x - end. - - - (* FIXME: Should be able to use the spec to avoid destructing here? *) - destruct c, chunk, addr, args; simplify; rt; simplify. - - + (** Preamble *) - invert MARR. simplify. - - unfold Op.eval_addressing in H0. - destruct (Archi.ptr64) eqn:ARCHI; simplify. - - unfold reg_stack_based_pointers in RSBP. - pose proof (RSBP r0) as RSBPr0. - - destruct (Registers.Regmap.get r0 rs) eqn:EQr0; simplify. - - rewrite ARCHI in H1. simplify. - subst. - - pose proof MASSOC as MASSOC'. - invert MASSOC'. - pose proof (H0 r0). - assert (HPler0 : Ple r0 (RTL.max_reg_function f)) - by (eapply RTL.max_reg_function_use; eauto; simplify; eauto). - apply H6 in HPler0. - invert HPler0; try congruence. - rewrite EQr0 in H8. - invert H8. - clear H0. clear H6. - - unfold check_address_parameter_signed in *; - unfold check_address_parameter_unsigned in *; simplify. - - remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0)) - (Integers.Ptrofs.of_int (Integers.Int.repr z))) as OFFSET. - - (** Modular preservation proof *) - assert (Integers.Ptrofs.unsigned OFFSET mod 4 = 0) as MOD_PRESERVE. - { rewrite HeqOFFSET. - apply PtrofsExtra.add_mod; simplify; try lia. - exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *) - rewrite Integers.Ptrofs.signed_repr; try assumption. - admit. (* FIXME: Register bounds. *) - apply PtrofsExtra.of_int_mod. - rewrite Integers.Int.signed_repr; simplify; try split; try assumption. - } - - (** Read bounds proof *) - assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH. - { destruct (Integers.Ptrofs.unsigned OFFSET destruct x eqn:EQ end; try reflexivity. - assert (Mem.valid_access m AST.Mint32 sp' - (Integers.Ptrofs.unsigned - (Integers.Ptrofs.add (Integers.Ptrofs.repr 0) - (Integers.Ptrofs.repr ptr))) Writable). - { pose proof H1. eapply Mem.store_valid_access_2 in H0. - exact H0. eapply Mem.store_valid_access_3. eassumption. } - pose proof (Mem.valid_access_store m AST.Mint32 sp' - (Integers.Ptrofs.unsigned - (Integers.Ptrofs.add (Integers.Ptrofs.repr 0) - (Integers.Ptrofs.repr ptr))) v). - apply X in H0. invert H0. congruence. - - + (** Preamble *) - invert MARR. simplify. - - unfold Op.eval_addressing in H0. - destruct (Archi.ptr64) eqn:ARCHI; simplify. - - unfold reg_stack_based_pointers in RSBP. - pose proof (RSBP r0) as RSBPr0. - pose proof (RSBP r1) as RSBPr1. - - destruct (Registers.Regmap.get r0 rs) eqn:EQr0; - destruct (Registers.Regmap.get r1 rs) eqn:EQr1; simplify. - - rewrite ARCHI in H1. simplify. - subst. - clear RSBPr1. - - pose proof MASSOC as MASSOC'. - invert MASSOC'. - pose proof (H0 r0). - pose proof (H0 r1). - assert (HPler0 : Ple r0 (RTL.max_reg_function f)) - by (eapply RTL.max_reg_function_use; eauto; simplify; eauto). - assert (HPler1 : Ple r1 (RTL.max_reg_function f)) - by (eapply RTL.max_reg_function_use; eauto; simpl; auto). - apply H6 in HPler0. - apply H8 in HPler1. - invert HPler0; invert HPler1; try congruence. - rewrite EQr0 in H10. - rewrite EQr1 in H12. - invert H10. invert H12. - clear H0. clear H6. clear H8. - - unfold check_address_parameter_signed in *; - unfold check_address_parameter_unsigned in *; simplify. - - remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0)) - (Integers.Ptrofs.of_int - (Integers.Int.add (Integers.Int.mul (valueToInt asr # r1) (Integers.Int.repr z)) - (Integers.Int.repr z0)))) as OFFSET. - - (** Modular preservation proof *) - assert (Integers.Ptrofs.unsigned OFFSET mod 4 = 0) as MOD_PRESERVE. - { rewrite HeqOFFSET. - apply PtrofsExtra.add_mod; simplify; try lia. - exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *) - rewrite Integers.Ptrofs.signed_repr; try assumption. - admit. (* FIXME: Register bounds. *) - apply PtrofsExtra.of_int_mod. - apply IntExtra.add_mod; simplify; try lia. - exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *) - apply IntExtra.mul_mod; simplify; try lia. - exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *) - admit. (* FIXME: Register bounds. *) - rewrite Integers.Int.signed_repr; simplify; try split; try assumption. - rewrite Integers.Int.signed_repr; simplify; try split; try assumption. - } - - (** Write bounds proof *) - assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as WRITE_BOUND_HIGH. - { destruct (Integers.Ptrofs.unsigned OFFSET destruct x eqn:EQ end; try reflexivity. - assert (Mem.valid_access m AST.Mint32 sp' - (Integers.Ptrofs.unsigned - (Integers.Ptrofs.add (Integers.Ptrofs.repr 0) - (Integers.Ptrofs.repr ptr))) Writable). - { pose proof H1. eapply Mem.store_valid_access_2 in H0. - exact H0. eapply Mem.store_valid_access_3. eassumption. } - pose proof (Mem.valid_access_store m AST.Mint32 sp' - (Integers.Ptrofs.unsigned - (Integers.Ptrofs.add (Integers.Ptrofs.repr 0) - (Integers.Ptrofs.repr ptr))) v). - apply X in H0. invert H0. congruence. - - + invert MARR. simplify. - - unfold Op.eval_addressing in H0. - destruct (Archi.ptr64) eqn:ARCHI; simplify. - rewrite ARCHI in H0. simplify. - - unfold check_address_parameter_unsigned in *; - unfold check_address_parameter_signed in *; simplify. - - assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity. - rewrite ZERO in H1. clear ZERO. - rewrite Integers.Ptrofs.add_zero_l in H1. - - remember i0 as OFFSET. - - (** Modular preservation proof *) - rename H0 into MOD_PRESERVE. - - (** Write bounds proof *) - assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as WRITE_BOUND_HIGH. - { destruct (Integers.Ptrofs.unsigned OFFSET destruct x eqn:EQ end; try reflexivity. - assert (Mem.valid_access m AST.Mint32 sp' - (Integers.Ptrofs.unsigned - (Integers.Ptrofs.add (Integers.Ptrofs.repr 0) - (Integers.Ptrofs.repr ptr))) Writable). - { pose proof H1. eapply Mem.store_valid_access_2 in H0. - exact H0. eapply Mem.store_valid_access_3. eassumption. } - pose proof (Mem.valid_access_store m AST.Mint32 sp' - (Integers.Ptrofs.unsigned - (Integers.Ptrofs.add (Integers.Ptrofs.repr 0) - (Integers.Ptrofs.repr ptr))) v). - apply X in H0. invert H0. congruence. - - - eexists. split. apply Smallstep.plus_one. + replace (Integers.Ptrofs.repr 0) with (Integers.Ptrofs.zero) by reflexivity. + rewrite Integers.Ptrofs.add_zero_l. + rewrite Integers.Ptrofs.unsigned_repr. intros. simplify. congruence. + apply Integers.Ptrofs.unsigned_range_2. } + + (** Normalisation proof *) + assert (Integers.Ptrofs.repr + (4 * Integers.Ptrofs.unsigned + (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4))) = OFFSET) + as NORMALISE. + { replace 4 with (Integers.Ptrofs.unsigned (Integers.Ptrofs.repr 4)) at 1 by reflexivity. + rewrite <- PtrofsExtra.mul_unsigned. + apply PtrofsExtra.mul_divu; simplify; auto; try lia. } + + (** Normalised bounds proof *) + assert (0 <= + Integers.Ptrofs.unsigned (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4)) + < (RTL.fn_stacksize f / 4)) + as NORMALISE_BOUND. + { split. + apply Integers.Ptrofs.unsigned_range_2. + assert (forall x y, Integers.Ptrofs.divu x y = Integers.Ptrofs.divu x y ) by reflexivity. + unfold Integers.Ptrofs.divu at 2 in H0. + rewrite H0. clear H0. + rewrite Integers.Ptrofs.unsigned_repr; simplify. + apply Zmult_lt_reg_r with (p := 4); try lia. + repeat rewrite ZLib.div_mul_undo; try lia. + split. + apply Z.div_pos; try lia; apply Integers.Ptrofs.unsigned_range_2. + apply Z.div_le_upper_bound; lia. } + + inversion NORMALISE_BOUND as [ NORMALISE_BOUND_LOW NORMALISE_BOUND_HIGH ]; + clear NORMALISE_BOUND. + + (** Start of proof proper *) + eexists. split. + eapply Smallstep.plus_one. eapply HTL.step_module; eauto. apply assumption_32bit. - eapply Verilog.stmnt_runp_Vnonblock_reg with - (rhsval := if b then posToValue 32 ifso else posToValue 32 ifnot). + econstructor. econstructor. econstructor. simplify. + econstructor. econstructor. econstructor. econstructor. simplify. - constructor. + all: simplify. - simpl. - destruct b. - eapply Verilog.erun_Vternary_true. - eapply eval_cond_correct; eauto. - constructor. - apply boolToValue_ValueToBool. - eapply Verilog.erun_Vternary_false. - eapply eval_cond_correct; eauto. - constructor. - apply boolToValue_ValueToBool. - constructor. + (** Verilog array lookup *) + unfold Verilog.arr_assocmap_lookup. setoid_rewrite H5. + f_equal. + + (** State Lookup *) unfold Verilog.merge_regs. + simplify. unfold_merge. + rewrite AssocMap.gso. apply AssocMap.gss. + apply st_greater_than_res. - destruct b. + (** Match states *) rewrite assumption_32bit. - simplify. - apply match_state with (sp' := sp'); eauto. - unfold Verilog.merge_regs. - unfold_merge. - apply regs_lessdef_add_greater. apply greater_than_max_func. + econstructor; eauto. + + (** Match assocmaps *) + unfold Verilog.merge_regs. simplify. unfold_merge. + apply regs_lessdef_add_match. + + (** Equality proof *) + assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity. + rewrite ZERO in H7. clear ZERO. + setoid_rewrite Integers.Ptrofs.add_zero_l in H7. + + specialize (H7 (Integers.Ptrofs.unsigned + (Integers.Ptrofs.divu + OFFSET + (Integers.Ptrofs.repr 4)))). + + exploit H7. + rewrite Z2Nat.id; eauto. + apply Z.div_pos; lia. + + intros I. + assert (Z.to_nat + (Integers.Ptrofs.unsigned + (Integers.Ptrofs.divu + OFFSET + (Integers.Ptrofs.repr 4))) + = + valueToNat (ZToValue 32 (Integers.Ptrofs.unsigned OFFSET / 4))) + as EXPR_OK by admit. + rewrite <- EXPR_OK. + rewrite NORMALISE in I. + rewrite H1 in I. + invert I. assumption. + + (** PC match *) + apply regs_lessdef_add_greater. + apply greater_than_max_func. assumption. - unfold state_st_wf. intros. - invert H3. - unfold Verilog.merge_regs. unfold_merge. + (** States well formed *) + unfold state_st_wf. inversion 1. simplify. + unfold Verilog.merge_regs. + unfold_merge. rewrite AssocMap.gso. apply AssocMap.gss. + apply st_greater_than_res. (** Match arrays *) - invert MARR. simplify. econstructor. repeat split; simplify. unfold HTL.empty_stack. @@ -2042,7 +1190,7 @@ Section CORRECTNESS. 2: { reflexivity. } rewrite AssocMap.gss. unfold Verilog.merge_arr. - setoid_rewrite H4. + setoid_rewrite H5. reflexivity. rewrite combine_length. @@ -2059,19 +1207,442 @@ Section CORRECTNESS. eauto. apply combine_none. assumption. + (** RSBP preservation *) + unfold reg_stack_based_pointers. intros. + destruct (Pos.eq_dec r0 dst); try rewrite e. (* FIXME: Prepare this for automation *) + + rewrite Registers.Regmap.gss. + unfold arr_stack_based_pointers in ASBP. + specialize (ASBP (Integers.Ptrofs.unsigned + (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4)))). + exploit ASBP; auto; intros I. + + rewrite NORMALISE in I. + assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity. + rewrite ZERO in I. clear ZERO. + simplify. + rewrite Integers.Ptrofs.add_zero_l in I. + rewrite H1 in I. + assumption. + simplify. + + rewrite Registers.Regmap.gso; auto. + Admitted. + Hint Resolve transl_iload_correct : htlproof. + + Lemma transl_istore_correct: + forall (s : list RTL.stackframe) (f : RTL.function) (sp : Values.val) (pc : positive) + (rs : Registers.Regmap.t Values.val) (m : mem) (chunk : AST.memory_chunk) + (addr : Op.addressing) (args : list Registers.reg) (src : Registers.reg) + (pc' : RTL.node) (a : Values.val) (m' : mem), + (RTL.fn_code f) ! pc = Some (RTL.Istore chunk addr args src pc') -> + Op.eval_addressing ge sp addr (map (fun r : positive => Registers.Regmap.get r rs) args) = Some a -> + Mem.storev chunk m a (Registers.Regmap.get src rs) = Some m' -> + forall R1 : HTL.state, + match_states (RTL.State s f sp pc rs m) R1 -> + exists R2 : HTL.state, + Smallstep.plus HTL.step tge R1 Events.E0 R2 /\ match_states (RTL.State s f sp pc' rs m') R2. + Proof. + intros s f sp pc rs m chunk addr args src pc' a m' H H0 H1 R1 MSTATES. + inv_state. + + destruct c, chunk, addr, args; simplify; tac; simplify. + + (** Preamble *) + invert MARR. simplify. + + unfold Op.eval_addressing in H0. + destruct (Archi.ptr64) eqn:ARCHI; simplify. + + unfold reg_stack_based_pointers in RSBP. + pose proof (RSBP r0) as RSBPr0. + + destruct (Registers.Regmap.get r0 rs) eqn:EQr0; simplify. + + rewrite ARCHI in H1. simplify. + subst. + + pose proof MASSOC as MASSOC'. + invert MASSOC'. + pose proof (H0 r0). + assert (HPler0 : Ple r0 (RTL.max_reg_function f)) + by (eapply RTL.max_reg_function_use; eauto; simplify; eauto). + apply H6 in HPler0. + invert HPler0; try congruence. + rewrite EQr0 in H8. + invert H8. + clear H0. clear H6. + + unfold check_address_parameter_unsigned in *; + unfold check_address_parameter_signed in *; simplify. + + remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0)) + (Integers.Ptrofs.of_int (Integers.Int.repr z))) as OFFSET. + + (** Modular preservation proof *) + assert (Integers.Ptrofs.unsigned OFFSET mod 4 = 0) as MOD_PRESERVE. + { rewrite HeqOFFSET. + apply PtrofsExtra.add_mod; simplify; try lia. + exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *) + rewrite Integers.Ptrofs.signed_repr; try assumption. + admit. (* FIXME: Register bounds. *) + apply PtrofsExtra.of_int_mod. + rewrite Integers.Int.signed_repr; simplify; try split; try assumption. + } + + (** Write bounds proof *) + assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as WRITE_BOUND_HIGH. + { destruct (Integers.Ptrofs.unsigned OFFSET destruct x eqn:EQ end; try reflexivity. + assert (Mem.valid_access m AST.Mint32 sp' + (Integers.Ptrofs.unsigned + (Integers.Ptrofs.add (Integers.Ptrofs.repr 0) + (Integers.Ptrofs.repr ptr))) Writable). + { pose proof H1. eapply Mem.store_valid_access_2 in H0. + exact H0. eapply Mem.store_valid_access_3. eassumption. } + pose proof (Mem.valid_access_store m AST.Mint32 sp' + (Integers.Ptrofs.unsigned + (Integers.Ptrofs.add (Integers.Ptrofs.repr 0) + (Integers.Ptrofs.repr ptr))) v). + apply X in H0. invert H0. congruence. + + + (** Preamble *) invert MARR. simplify. + + unfold Op.eval_addressing in H0. + destruct (Archi.ptr64) eqn:ARCHI; simplify. + + unfold reg_stack_based_pointers in RSBP. + pose proof (RSBP r0) as RSBPr0. + pose proof (RSBP r1) as RSBPr1. + + destruct (Registers.Regmap.get r0 rs) eqn:EQr0; + destruct (Registers.Regmap.get r1 rs) eqn:EQr1; simplify. + + rewrite ARCHI in H1. simplify. + subst. + clear RSBPr1. + + pose proof MASSOC as MASSOC'. + invert MASSOC'. + pose proof (H0 r0). + pose proof (H0 r1). + assert (HPler0 : Ple r0 (RTL.max_reg_function f)) + by (eapply RTL.max_reg_function_use; eauto; simplify; eauto). + assert (HPler1 : Ple r1 (RTL.max_reg_function f)) + by (eapply RTL.max_reg_function_use; eauto; simpl; auto). + apply H6 in HPler0. + apply H8 in HPler1. + invert HPler0; invert HPler1; try congruence. + rewrite EQr0 in H10. + rewrite EQr1 in H12. + invert H10. invert H12. + clear H0. clear H6. clear H8. + + unfold check_address_parameter_signed in *; + unfold check_address_parameter_unsigned in *; simplify. + + remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0)) + (Integers.Ptrofs.of_int + (Integers.Int.add (Integers.Int.mul (valueToInt asr # r1) (Integers.Int.repr z)) + (Integers.Int.repr z0)))) as OFFSET. + + (** Modular preservation proof *) + assert (Integers.Ptrofs.unsigned OFFSET mod 4 = 0) as MOD_PRESERVE. + { rewrite HeqOFFSET. + apply PtrofsExtra.add_mod; simplify; try lia. + exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *) + rewrite Integers.Ptrofs.signed_repr; try assumption. + admit. (* FIXME: Register bounds. *) + apply PtrofsExtra.of_int_mod. + apply IntExtra.add_mod; simplify; try lia. + exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *) + apply IntExtra.mul_mod; simplify; try lia. + exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *) + admit. (* FIXME: Register bounds. *) + rewrite Integers.Int.signed_repr; simplify; try split; try assumption. + rewrite Integers.Int.signed_repr; simplify; try split; try assumption. + } + + (** Write bounds proof *) + assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as WRITE_BOUND_HIGH. + { destruct (Integers.Ptrofs.unsigned OFFSET destruct x eqn:EQ end; try reflexivity. + assert (Mem.valid_access m AST.Mint32 sp' + (Integers.Ptrofs.unsigned + (Integers.Ptrofs.add (Integers.Ptrofs.repr 0) + (Integers.Ptrofs.repr ptr))) Writable). + { pose proof H1. eapply Mem.store_valid_access_2 in H0. + exact H0. eapply Mem.store_valid_access_3. eassumption. } + pose proof (Mem.valid_access_store m AST.Mint32 sp' + (Integers.Ptrofs.unsigned + (Integers.Ptrofs.add (Integers.Ptrofs.repr 0) + (Integers.Ptrofs.repr ptr))) v). + apply X in H0. invert H0. congruence. + + + invert MARR. simplify. + + unfold Op.eval_addressing in H0. + destruct (Archi.ptr64) eqn:ARCHI; simplify. + rewrite ARCHI in H0. simplify. + + unfold check_address_parameter_unsigned in *; + unfold check_address_parameter_signed in *; simplify. + + assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity. + rewrite ZERO in H1. clear ZERO. + rewrite Integers.Ptrofs.add_zero_l in H1. + + remember i0 as OFFSET. + + (** Modular preservation proof *) + rename H0 into MOD_PRESERVE. + + (** Write bounds proof *) + assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as WRITE_BOUND_HIGH. + { destruct (Integers.Ptrofs.unsigned OFFSET destruct x eqn:EQ end; try reflexivity. + assert (Mem.valid_access m AST.Mint32 sp' + (Integers.Ptrofs.unsigned + (Integers.Ptrofs.add (Integers.Ptrofs.repr 0) + (Integers.Ptrofs.repr ptr))) Writable). + { pose proof H1. eapply Mem.store_valid_access_2 in H0. + exact H0. eapply Mem.store_valid_access_3. eassumption. } + pose proof (Mem.valid_access_store m AST.Mint32 sp' + (Integers.Ptrofs.unsigned + (Integers.Ptrofs.add (Integers.Ptrofs.repr 0) + (Integers.Ptrofs.repr ptr))) v). + apply X in H0. invert H0. congruence. + Admitted. + Hint Resolve transl_istore_correct : htlproof. + + Lemma transl_icond_correct: + forall (s : list RTL.stackframe) (f : RTL.function) (sp : Values.val) (pc : positive) + (rs : Registers.Regmap.t Values.val) (m : mem) (cond : Op.condition) (args : list Registers.reg) + (ifso ifnot : RTL.node) (b : bool) (pc' : RTL.node), + (RTL.fn_code f) ! pc = Some (RTL.Icond cond args ifso ifnot) -> + Op.eval_condition cond (map (fun r : positive => Registers.Regmap.get r rs) args) m = Some b -> + pc' = (if b then ifso else ifnot) -> + forall R1 : HTL.state, + match_states (RTL.State s f sp pc rs m) R1 -> + exists R2 : HTL.state, + Smallstep.plus HTL.step tge R1 Events.E0 R2 /\ match_states (RTL.State s f sp pc' rs m) R2. + Proof. + intros s f sp pc rs m cond args ifso ifnot b pc' H H0 H1 R1 MSTATE. + inv_state. + + eexists. split. apply Smallstep.plus_one. + eapply HTL.step_module; eauto. + apply assumption_32bit. + eapply Verilog.stmnt_runp_Vnonblock_reg with + (rhsval := if b then posToValue 32 ifso else posToValue 32 ifnot). - - (* Return *) - econstructor. split. + constructor. + + simpl. + destruct b. + eapply Verilog.erun_Vternary_true. + eapply eval_cond_correct; eauto. + constructor. + apply boolToValue_ValueToBool. + eapply Verilog.erun_Vternary_false. + eapply eval_cond_correct; eauto. + constructor. + apply boolToValue_ValueToBool. + constructor. + unfold Verilog.merge_regs. + unfold_merge. + apply AssocMap.gss. + + destruct b. + rewrite assumption_32bit. + simplify. + apply match_state with (sp' := sp'); eauto. + unfold Verilog.merge_regs. + unfold_merge. + apply regs_lessdef_add_greater. apply greater_than_max_func. + assumption. + + unfold state_st_wf. intros. + invert H3. + unfold Verilog.merge_regs. unfold_merge. + apply AssocMap.gss. + + (** Match arrays *) + invert MARR. simplify. + econstructor. + repeat split; simplify. + unfold HTL.empty_stack. + simplify. + unfold Verilog.merge_arrs. + + rewrite AssocMap.gcombine. + 2: { reflexivity. } + rewrite AssocMap.gss. + unfold Verilog.merge_arr. + setoid_rewrite H4. + reflexivity. + + rewrite combine_length. + unfold arr_repeat. simplify. + rewrite list_repeat_len. + reflexivity. + + unfold arr_repeat. simplify. + rewrite list_repeat_len. + congruence. + + intros. + erewrite array_get_error_equal. + eauto. apply combine_none. + assumption. + + rewrite assumption_32bit. + apply match_state with (sp' := sp'); eauto. + unfold Verilog.merge_regs. unfold_merge. + apply regs_lessdef_add_greater. apply greater_than_max_func. + assumption. + + unfold state_st_wf. intros. + invert H1. + unfold Verilog.merge_regs. unfold_merge. + apply AssocMap.gss. + + (** Match arrays *) + invert MARR. simplify. + econstructor. + repeat split; simplify. + unfold HTL.empty_stack. + simplify. + unfold Verilog.merge_arrs. + + rewrite AssocMap.gcombine. + 2: { reflexivity. } + rewrite AssocMap.gss. + unfold Verilog.merge_arr. + setoid_rewrite H2. + reflexivity. + + rewrite combine_length. + unfold arr_repeat. simplify. + rewrite list_repeat_len. + reflexivity. + + unfold arr_repeat. simplify. + rewrite list_repeat_len. + congruence. + + intros. + erewrite array_get_error_equal. + eauto. apply combine_none. + assumption. + + Unshelve. + constructor. + Qed. + Hint Resolve transl_icond_correct : htlproof. + + Lemma transl_ijumptable_correct: + forall (s : list RTL.stackframe) (f : RTL.function) (sp : Values.val) (pc : positive) + (rs : Registers.Regmap.t Values.val) (m : mem) (arg : Registers.reg) (tbl : list RTL.node) + (n : Integers.Int.int) (pc' : RTL.node), + (RTL.fn_code f) ! pc = Some (RTL.Ijumptable arg tbl) -> + Registers.Regmap.get arg rs = Values.Vint n -> + list_nth_z tbl (Integers.Int.unsigned n) = Some pc' -> + forall R1 : HTL.state, + match_states (RTL.State s f sp pc rs m) R1 -> + exists R2 : HTL.state, + Smallstep.plus HTL.step tge R1 Events.E0 R2 /\ match_states (RTL.State s f sp pc' rs m) R2. + Proof. + intros s f sp pc rs m arg tbl n pc' H H0 H1 R1 MSTATE. + Admitted. + Hint Resolve transl_ijumptable_correct : htlproof. + + Lemma transl_ireturn_correct: + forall (s : list RTL.stackframe) (f : RTL.function) (stk : Values.block) + (pc : positive) (rs : RTL.regset) (m : mem) (or : option Registers.reg) + (m' : mem), + (RTL.fn_code f) ! pc = Some (RTL.Ireturn or) -> + Mem.free m stk 0 (RTL.fn_stacksize f) = Some m' -> + forall R1 : HTL.state, + match_states (RTL.State s f (Values.Vptr stk Integers.Ptrofs.zero) pc rs m) R1 -> + exists R2 : HTL.state, + Smallstep.plus HTL.step tge R1 Events.E0 R2 /\ + match_states (RTL.Returnstate s (Registers.regmap_optget or Values.Vundef rs) m') R2. + Proof. + intros s f stk pc rs m or m' H H0 R1 MSTATE. + inv_state. + + - econstructor. split. eapply Smallstep.plus_two. - + eapply HTL.step_module; eauto. apply assumption_32bit. constructor. @@ -2135,6 +2241,7 @@ Section CORRECTNESS. constructor; auto. constructor. + (* FIXME: Duplication *) - econstructor. split. eapply Smallstep.plus_two. eapply HTL.step_module; eauto. @@ -2171,104 +2278,134 @@ Section CORRECTNESS. apply H1. eapply RTL.max_reg_function_use. eauto. simpl; tauto. apply st_greater_than_res. - - inversion MSTATE; subst. inversion TF; subst. - econstructor. split. apply Smallstep.plus_one. - eapply HTL.step_call. simplify. + Unshelve. + all: constructor. + Qed. + Hint Resolve transl_ireturn_correct : htlproof. + + Lemma transl_callstate_correct: + forall (s : list RTL.stackframe) (f : RTL.function) (args : list Values.val) + (m : mem) (m' : Mem.mem') (stk : Values.block), + Mem.alloc m 0 (RTL.fn_stacksize f) = (m', stk) -> + forall R1 : HTL.state, + match_states (RTL.Callstate s (AST.Internal f) args m) R1 -> + exists R2 : HTL.state, + Smallstep.plus HTL.step tge R1 Events.E0 R2 /\ + match_states + (RTL.State s f (Values.Vptr stk Integers.Ptrofs.zero) (RTL.fn_entrypoint f) + (RTL.init_regs args (RTL.fn_params f)) m') R2. + Proof. + intros s f args m m' stk H R1 MSTATE. - apply match_state with (sp' := stk); eauto. + inversion MSTATE; subst. inversion TF; subst. + econstructor. split. apply Smallstep.plus_one. + eapply HTL.step_call. simplify. - apply regs_lessdef_add_greater. - apply greater_than_max_func. - apply init_reg_assoc_empty. - unfold state_st_wf. - intros. inv H3. apply AssocMap.gss. + apply match_state with (sp' := stk); eauto. - constructor. - - econstructor. simplify. - repeat split. unfold HTL.empty_stack. - simplify. apply AssocMap.gss. - unfold arr_repeat. simplify. - apply list_repeat_len. - intros. - destruct (Mem.load AST.Mint32 m' stk - (Integers.Ptrofs.unsigned (Integers.Ptrofs.add - Integers.Ptrofs.zero - (Integers.Ptrofs.repr (4 * ptr))))) eqn:LOAD. - pose proof Mem.load_alloc_same as LOAD_ALLOC. - pose proof H as ALLOC. - eapply LOAD_ALLOC in ALLOC. - 2: { exact LOAD. } - rewrite ALLOC. - repeat constructor. - constructor. + apply regs_lessdef_add_greater. + apply greater_than_max_func. + apply init_reg_assoc_empty. + unfold state_st_wf. + intros. inv H3. apply AssocMap.gss. - unfold reg_stack_based_pointers. intros. - unfold RTL.init_regs; simplify. - destruct (RTL.fn_params f); - rewrite Registers.Regmap.gi; constructor. + constructor. - unfold arr_stack_based_pointers. intros. - simplify. - destruct (Mem.load AST.Mint32 m' stk - (Integers.Ptrofs.unsigned (Integers.Ptrofs.add - Integers.Ptrofs.zero - (Integers.Ptrofs.repr (4 * ptr))))) eqn:LOAD. - pose proof Mem.load_alloc_same as LOAD_ALLOC. - pose proof H as ALLOC. - eapply LOAD_ALLOC in ALLOC. - 2: { exact LOAD. } - rewrite ALLOC. - repeat constructor. - constructor. + econstructor. simplify. + repeat split. unfold HTL.empty_stack. + simplify. apply AssocMap.gss. + unfold arr_repeat. simplify. + apply list_repeat_len. + intros. + destruct (Mem.load AST.Mint32 m' stk + (Integers.Ptrofs.unsigned (Integers.Ptrofs.add + Integers.Ptrofs.zero + (Integers.Ptrofs.repr (4 * ptr))))) eqn:LOAD. + pose proof Mem.load_alloc_same as LOAD_ALLOC. + pose proof H as ALLOC. + eapply LOAD_ALLOC in ALLOC. + 2: { exact LOAD. } + rewrite ALLOC. + repeat constructor. + constructor. - Transparent Mem.alloc. (* TODO: Since there are opaque there's probably a lemma. *) - Transparent Mem.load. - Transparent Mem.store. - unfold stack_bounds. - split. + unfold reg_stack_based_pointers. intros. + unfold RTL.init_regs; simplify. + destruct (RTL.fn_params f); + rewrite Registers.Regmap.gi; constructor. + + unfold arr_stack_based_pointers. intros. + simplify. + destruct (Mem.load AST.Mint32 m' stk + (Integers.Ptrofs.unsigned (Integers.Ptrofs.add + Integers.Ptrofs.zero + (Integers.Ptrofs.repr (4 * ptr))))) eqn:LOAD. + pose proof Mem.load_alloc_same as LOAD_ALLOC. + pose proof H as ALLOC. + eapply LOAD_ALLOC in ALLOC. + 2: { exact LOAD. } + rewrite ALLOC. + repeat constructor. + constructor. - unfold Mem.alloc in H. - invert H. - simplify. - unfold Mem.load. - intros. - match goal with | |- context[if ?x then _ else _] => destruct x end; try congruence. - invert v0. unfold Mem.range_perm in H3. - unfold Mem.perm in H3. simplify. - unfold Mem.perm_order' in H3. - rewrite Integers.Ptrofs.add_zero_l in H3. - rewrite Integers.Ptrofs.unsigned_repr in H3; simplify; try lia. - exploit (H3 ptr). lia. intros. - rewrite Maps.PMap.gss in H8. - match goal with | H8 : context[if ?x then _ else _] |- _ => destruct x eqn:EQ end; try contradiction. - simplify. - apply proj_sumbool_true in H10. lia. + Transparent Mem.alloc. (* TODO: Since there are opaque there's probably a lemma. *) + Transparent Mem.load. + Transparent Mem.store. + unfold stack_bounds. + split. - unfold Mem.alloc in H. - invert H. - simplify. - unfold Mem.store. - intros. - match goal with | |- context[if ?x then _ else _] => destruct x end; try congruence. - invert v0. unfold Mem.range_perm in H3. - unfold Mem.perm in H3. simplify. - unfold Mem.perm_order' in H3. - rewrite Integers.Ptrofs.add_zero_l in H3. - rewrite Integers.Ptrofs.unsigned_repr in H3; simplify; try lia. - exploit (H3 ptr). lia. intros. - rewrite Maps.PMap.gss in H8. - match goal with | H8 : context[if ?x then _ else _] |- _ => destruct x eqn:EQ end; try contradiction. - simplify. - apply proj_sumbool_true in H10. lia. - Opaque Mem.alloc. - Opaque Mem.load. - Opaque Mem.store. - - - inversion MSTATE. - inversion MF. - Admitted. - Hint Resolve transl_step_correct : htlproof. + unfold Mem.alloc in H. + invert H. + simplify. + unfold Mem.load. + intros. + match goal with | |- context[if ?x then _ else _] => destruct x end; try congruence. + invert v0. unfold Mem.range_perm in H3. + unfold Mem.perm in H3. simplify. + unfold Mem.perm_order' in H3. + rewrite Integers.Ptrofs.add_zero_l in H3. + rewrite Integers.Ptrofs.unsigned_repr in H3; simplify; try lia. + exploit (H3 ptr). lia. intros. + rewrite Maps.PMap.gss in H8. + match goal with | H8 : context[if ?x then _ else _] |- _ => destruct x eqn:EQ end; try contradiction. + simplify. + apply proj_sumbool_true in H10. lia. + + unfold Mem.alloc in H. + invert H. + simplify. + unfold Mem.store. + intros. + match goal with | |- context[if ?x then _ else _] => destruct x end; try congruence. + invert v0. unfold Mem.range_perm in H3. + unfold Mem.perm in H3. simplify. + unfold Mem.perm_order' in H3. + rewrite Integers.Ptrofs.add_zero_l in H3. + rewrite Integers.Ptrofs.unsigned_repr in H3; simplify; try lia. + exploit (H3 ptr). lia. intros. + rewrite Maps.PMap.gss in H8. + match goal with | H8 : context[if ?x then _ else _] |- _ => destruct x eqn:EQ end; try contradiction. + simplify. + apply proj_sumbool_true in H10. lia. + Opaque Mem.alloc. + Opaque Mem.load. + Opaque Mem.store. + Qed. + Hint Resolve transl_callstate_correct : htlproof. + + Lemma transl_returnstate_correct: + forall (res0 : Registers.reg) (f : RTL.function) (sp : Values.val) (pc : RTL.node) + (rs : RTL.regset) (s : list RTL.stackframe) (vres : Values.val) (m : mem) + (R1 : HTL.state), + match_states (RTL.Returnstate (RTL.Stackframe res0 f sp pc rs :: s) vres m) R1 -> + exists R2 : HTL.state, + Smallstep.plus HTL.step tge R1 Events.E0 R2 /\ + match_states (RTL.State s f sp pc (Registers.Regmap.set res0 vres rs) m) R2. + Proof. + intros res0 f sp pc rs s vres m R1 MSTATE. + inversion MSTATE. inversion MF. + Qed. + Hint Resolve transl_returnstate_correct : htlproof. Lemma option_inv : forall A x y, @@ -2292,7 +2429,6 @@ Section CORRECTNESS. trivial. symmetry; eapply Linking.match_program_main; eauto. Qed. - (* Had to admit proof because currently there is no way to force main to be Internal. *) Lemma transl_initial_states : forall s1 : Smallstep.state (RTL.semantics prog), @@ -2344,14 +2480,22 @@ Section CORRECTNESS. Qed. Hint Resolve transl_final_states : htlproof. -Theorem transf_program_correct: - Smallstep.forward_simulation (RTL.semantics prog) (HTL.semantics tprog). -Proof. - eapply Smallstep.forward_simulation_plus. - apply senv_preserved. - eexact transl_initial_states. - eexact transl_final_states. - exact transl_step_correct. -Qed. + Theorem transl_step_correct: + forall (S1 : RTL.state) t S2, + RTL.step ge S1 t S2 -> + forall (R1 : HTL.state), + match_states S1 R1 -> + exists R2, Smallstep.plus HTL.step tge R1 t R2 /\ match_states S2 R2. + Proof. + induction 1; eauto with htlproof; (intros; inv_state). + Qed. + Hint Resolve transl_step_correct : htlproof. + + Theorem transf_program_correct: + Smallstep.forward_simulation (RTL.semantics prog) (HTL.semantics tprog). + Proof. + eapply Smallstep.forward_simulation_plus; eauto with htlproof. + apply senv_preserved. + Qed. End CORRECTNESS. -- cgit From f8ff27915f8c4d5fb6f31ec2a0a73f65cf604c43 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Tue, 30 Jun 2020 19:52:59 +0100 Subject: Add htl pretty printing --- driver/CoqupDriver.ml | 4 ++++ src/Compiler.v | 2 ++ src/CoqupClflags.ml | 1 + src/extraction/Extraction.v | 1 + src/verilog/PrintVerilog.ml | 2 +- src/verilog/PrintVerilog.mli | 2 ++ test/array.c | 7 +++++-- test/matrix.c | 16 +++++++--------- 8 files changed, 23 insertions(+), 12 deletions(-) diff --git a/driver/CoqupDriver.ml b/driver/CoqupDriver.ml index afbe4d0..0b64f48 100644 --- a/driver/CoqupDriver.ml +++ b/driver/CoqupDriver.ml @@ -65,6 +65,7 @@ let compile_c_file sourcename ifile ofile = set_dest Coqup.PrintClight.destination option_dclight ".light.c"; set_dest Coqup.PrintCminor.destination option_dcminor ".cm"; set_dest Coqup.PrintRTL.destination option_drtl ".rtl"; + set_dest Coqup.PrintHTL.destination option_dhtl ".htl"; set_dest Coqup.Regalloc.destination_alloctrace option_dalloctrace ".alloctrace"; set_dest Coqup.PrintLTL.destination option_dltl ".ltl"; set_dest Coqup.PrintMach.destination option_dmach ".mach"; @@ -253,6 +254,7 @@ Code generation options: (use -fno- to turn off -f) -dclight Save generated Clight in .light.c -dcminor Save generated Cminor in .cm -drtl Save RTL at various optimization points in .rtl. + -dhtl Save HTL before Verilog generation .htl -dltl Save LTL after register allocation in .ltl -dmach Save generated Mach code in .mach -dasm Save generated assembly in .s @@ -377,6 +379,7 @@ let cmdline_actions = Exact "-dclight", Set option_dclight; Exact "-dcminor", Set option_dcminor; Exact "-drtl", Set option_drtl; + Exact "-dhtl", Set option_dhtl; Exact "-dltl", Set option_dltl; Exact "-dalloctrace", Set option_dalloctrace; Exact "-dmach", Set option_dmach; @@ -388,6 +391,7 @@ let cmdline_actions = option_dclight := true; option_dcminor := true; option_drtl := true; + option_dhtl := true; option_dltl := true; option_dalloctrace := true; option_dmach := true; diff --git a/src/Compiler.v b/src/Compiler.v index 98ef429..a34b359 100644 --- a/src/Compiler.v +++ b/src/Compiler.v @@ -51,6 +51,7 @@ From coqup Require HTLgen. Parameter print_RTL: Z -> RTL.program -> unit. +Parameter print_HTL: HTL.program -> unit. Definition print {A: Type} (printer: A -> unit) (prog: A) : A := let unused := printer prog in prog. @@ -79,6 +80,7 @@ Definition transf_backend (r : RTL.program) : res Verilog.program := @@@ Inlining.transf_program @@ print (print_RTL 1) @@@ HTLgen.transl_program + @@ print print_HTL @@ Veriloggen.transl_program. Definition transf_frontend (p: Csyntax.program) : res RTL.program := diff --git a/src/CoqupClflags.ml b/src/CoqupClflags.ml index 83dd31d..5b40ce6 100644 --- a/src/CoqupClflags.ml +++ b/src/CoqupClflags.ml @@ -3,3 +3,4 @@ let option_simulate = ref false let option_hls = ref true let option_debug_hls = ref false let option_initial = ref false +let option_dhtl = ref false diff --git a/src/extraction/Extraction.v b/src/extraction/Extraction.v index ba87af6..df21dc4 100644 --- a/src/extraction/Extraction.v +++ b/src/extraction/Extraction.v @@ -128,6 +128,7 @@ Extract Constant Compiler.print_Clight => "PrintClight.print_if". Extract Constant Compiler.print_Cminor => "PrintCminor.print_if". Extract Constant driver.Compiler.print_RTL => "PrintRTL.print_if". Extract Constant Compiler.print_RTL => "PrintRTL.print_if". +Extract Constant Compiler.print_HTL => "PrintHTL.print_if". Extract Constant Compiler.print_LTL => "PrintLTL.print_if". Extract Constant Compiler.print_Mach => "PrintMach.print_if". Extract Constant Compiler.print => "fun (f: 'a -> unit) (x: 'a) -> f x; x". diff --git a/src/verilog/PrintVerilog.ml b/src/verilog/PrintVerilog.ml index 6d10887..5265c97 100644 --- a/src/verilog/PrintVerilog.ml +++ b/src/verilog/PrintVerilog.ml @@ -78,7 +78,7 @@ let rec pprint_expr = function | Vvari (s, i) -> concat [register s; "["; pprint_expr i; "]"] | Vinputvar s -> register s | Vunop (u, e) -> concat ["("; unop u; pprint_expr e; ")"] - | Vbinop (op, a, b) -> concat ["("; pprint_binop (pprint_expr a) (pprint_expr b) op; ")"] + | Vbinop (op, a, b) -> concat [pprint_binop (pprint_expr a) (pprint_expr b) op] | Vternary (c, t, f) -> concat ["("; pprint_expr c; " ? "; pprint_expr t; " : "; pprint_expr f; ")"] let rec pprint_stmnt i = diff --git a/src/verilog/PrintVerilog.mli b/src/verilog/PrintVerilog.mli index 6544e52..62bf63f 100644 --- a/src/verilog/PrintVerilog.mli +++ b/src/verilog/PrintVerilog.mli @@ -16,6 +16,8 @@ * along with this program. If not, see . *) +val pprint_stmnt : int -> Verilog.stmnt -> string + val print_value : out_channel -> Value.value -> unit val print_program : bool -> out_channel -> Verilog.program -> unit diff --git a/test/array.c b/test/array.c index e33d47b..7d78a61 100644 --- a/test/array.c +++ b/test/array.c @@ -1,4 +1,7 @@ int main() { - int x[5] = {1, 2, 3, 4, 5}; - return x[2]; + int x[3] = {1, 2, 3}; + int sum = 0, incr = 1; + for (int i = 0; i < 3; i=i+incr) + sum += x[i]; + return sum; } diff --git a/test/matrix.c b/test/matrix.c index 2bb17e7..daa00ae 100644 --- a/test/matrix.c +++ b/test/matrix.c @@ -1,10 +1,8 @@ -#define N 4 - -void matrix_multiply(int first[N][N], int second[N][N], int multiply[N][N]) { +void matrix_multiply(int first[2][2], int second[2][2], int multiply[2][2]) { int sum = 0; - for (int c = 0; c < N; c++) { - for (int d = 0; d < N; d++) { - for (int k = 0; k < N; k++) { + for (int c = 0; c < 2; c++) { + for (int d = 0; d < 2; d++) { + for (int k = 0; k < 2; k++) { sum = sum + first[c][k]*second[k][d]; } multiply[c][d] = sum; @@ -14,9 +12,9 @@ void matrix_multiply(int first[N][N], int second[N][N], int multiply[N][N]) { } int main() { - int f[N][N] = {{1, 2, 3, 4}, {1, 2, 3, 4}, {1, 2, 3, 4}, {1, 2, 3, 4}}; - int s[N][N] = {{5, 6, 7, 8}, {5, 6, 7, 8}, {5, 6, 7, 8}, {5, 6, 7, 8}}; - int m[N][N] = {{0, 0, 0, 0}, {0, 0, 0, 0}, {0, 0, 0, 0}, {0, 0, 0, 0}}; + int f[2][2] = {{1, 2}, {3, 4}}; + int s[2][2] = {{5, 6}, {7, 8}}; + int m[2][2] = {{0, 0}, {0, 0}}; matrix_multiply(f, s, m); return m[1][1]; -- cgit From fae3efb7755ae7d4832b85fb9244b3a93fa5c9eb Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Tue, 30 Jun 2020 19:53:24 +0100 Subject: Add HTL pretty printer --- src/verilog/PrintHTL.ml | 81 +++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 81 insertions(+) create mode 100644 src/verilog/PrintHTL.ml diff --git a/src/verilog/PrintHTL.ml b/src/verilog/PrintHTL.ml new file mode 100644 index 0000000..0bdba51 --- /dev/null +++ b/src/verilog/PrintHTL.ml @@ -0,0 +1,81 @@ +(* -*- mode: tuareg -*- + * CoqUp: Verified high-level synthesis. + * Copyright (C) 2019-2020 Yann Herklotz + * + * This program is free software: you can redistribute it and/or modify + * it under the terms of the GNU General Public License as published by + * the Free Software Foundation, either version 3 of the License, or + * (at your option) any later version. + * + * This program is distributed in the hope that it will be useful, + * but WITHOUT ANY WARRANTY; without even the implied warranty of + * MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the + * GNU General Public License for more details. + * + * You should have received a copy of the GNU General Public License + * along with this program. If not, see . + *) + +open Value +open Datatypes +open Camlcoq +open AST +open Clflags +open Printf +open Maps +open AST +open HTL +open PrintAST +open PrintVerilog +open CoqupClflags + +let pstr pp = fprintf pp "%s" + +let reg pp r = + fprintf pp "x%d" (P.to_int r) + +let rec regs pp = function + | [] -> () + | [r] -> reg pp r + | r1::rl -> fprintf pp "%a, %a" reg r1 regs rl + +let print_instruction pp (pc, i) = + fprintf pp "%5d:\t%s" pc (pprint_stmnt 0 i) + +let print_module pp id f = + fprintf pp "%s(%a) {\n" (extern_atom id) regs f.mod_params; + let datapath = + List.sort + (fun (pc1, _) (pc2, _) -> compare pc2 pc1) + (List.rev_map + (fun (pc, i) -> (P.to_int pc, i)) + (PTree.elements f.mod_datapath)) in + let controllogic = + List.sort + (fun (pc1, _) (pc2, _) -> compare pc2 pc1) + (List.rev_map + (fun (pc, i) -> (P.to_int pc, i)) + (PTree.elements f.mod_controllogic)) in + fprintf pp " datapath {\n"; + List.iter (print_instruction pp) datapath; + fprintf pp " }\n\n controllogic {\n"; + List.iter (print_instruction pp) controllogic; + fprintf pp " }\n}\n\n" + +let print_globdef pp (id, gd) = + match gd with + | Gfun(Internal f) -> print_module pp id f + | _ -> () + +let print_program pp prog = + List.iter (print_globdef pp) prog.prog_defs + +let destination : string option ref = ref None + +let print_if prog = + match !destination with + | None -> () + | Some f -> + let oc = open_out f in + print_program oc prog; + close_out oc -- cgit From f02b7b9a3879781ae332e4a967f605d961210000 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Tue, 30 Jun 2020 20:18:18 +0100 Subject: Heavy automation of proofs. --- src/common/Coquplib.v | 9 +- src/translation/HTLgenproof.v | 385 +++++++++--------------------------------- 2 files changed, 87 insertions(+), 307 deletions(-) diff --git a/src/common/Coquplib.v b/src/common/Coquplib.v index 5de1e7c..ba0a5dc 100644 --- a/src/common/Coquplib.v +++ b/src/common/Coquplib.v @@ -51,6 +51,13 @@ Ltac clear_obvious := | [ H : _ /\ _ |- _ ] => invert H end. +Ltac nicify_goals := + repeat match goal with + | [ |- _ /\ _ ] => split + | [ |- Some _ = Some _ ] => try reflexivity + | [ _ : ?x |- ?x ] => assumption + end. + Ltac kill_bools := repeat match goal with | [ H : _ && _ = true |- _ ] => apply andb_prop in H @@ -118,7 +125,7 @@ Ltac unfold_constants := end. Ltac simplify := unfold_constants; simpl in *; - repeat (clear_obvious; kill_bools); + repeat (clear_obvious; nicify_goals; kill_bools); simpl in *; try discriminate. Global Opaque Nat.div. diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index 252119a..9f62bb9 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -335,6 +335,17 @@ Proof. constructor. Qed. +Lemma arr_lookup_some: + forall (z : Z) (r0 : Registers.reg) (r : Verilog.reg) (asr : assocmap) (asa : Verilog.assocmap_arr) + (stack : Array (option value)) (H5 : asa ! r = Some stack) n, + exists x, Verilog.arr_assocmap_lookup asa r n = Some x. +Proof. + intros z r0 r asr asa stack H5 n. + eexists. + unfold Verilog.arr_assocmap_lookup. rewrite H5. reflexivity. +Qed. +Hint Resolve arr_lookup_some : htlproof. + Section CORRECTNESS. Variable prog : RTL.program. @@ -421,6 +432,42 @@ Section CORRECTNESS. exists asr' asa', HTL.step tge (HTL.State res m st asr asa) Events.E0 (HTL.State res m st asr' asa'). + Ltac big_tac := + repeat (simplify; + match goal with + | [ |- context[Verilog.merge_regs _ _] ] => + unfold Verilog.merge_regs; simplify; unfold_merge + | [ |- context[_ # ?d <- _ ! ?d] ] => apply AssocMap.gss + | [ |- context[_ # ?d <- _ ! ?s] ] => rewrite AssocMap.gso; try apply st_greater_than_res + | [ |- context[valueToPos (posToValue 32 _)] ] => rewrite assumption_32bit + | [ |- context[match_states _ _] ] => econstructor; eauto + | [ |- context[Verilog.merge_arr] ] => unfold Verilog.merge_arr; simplify + | [ |- context[(AssocMap.empty _) ! _] ] => rewrite AssocMap.gempty; simplify + + | [ H : ?asa ! ?r = Some _ |- Verilog.arr_assocmap_lookup ?asa ?r _ = Some _ ] => + unfold Verilog.arr_assocmap_lookup; setoid_rewrite H; f_equal + + | [ |- match_assocmaps _ _ _ # _ <- (posToValue 32 _) ] => + apply regs_lessdef_add_greater; [> apply greater_than_max_func | assumption] + + | [ |- state_st_wf _ _ ] => unfold state_st_wf; inversion 1; simplify + + | [ |- match_arrs _ _ _ _ _ ] => econstructor; simplify + | [ |- context[HTL.empty_stack] ] => unfold HTL.empty_stack; simplify + | [ |- context[Verilog.merge_arrs _ _] ] => unfold Verilog.merge_arrs; simplify + | [ |- context[(AssocMap.combine _ _ _) ! _] ] => + try (rewrite AssocMap.gcombine; [> | reflexivity]) + + | [ |- context[reg_stack_based_pointers] ] => unfold reg_stack_based_pointers; intros + | [ |- context[Registers.Regmap.get ?d (Registers.Regmap.set ?d _ _)] ] => + rewrite Registers.Regmap.gss + | [ |- context[Registers.Regmap.get ?s (Registers.Regmap.set ?d _ _)] ] => + destruct (Pos.eq_dec s d) as [EQ|EQ]; + [> rewrite EQ | rewrite Registers.Regmap.gso; auto] + + | [ H : _ ! _ = Some _ |- _] => try (setoid_rewrite H; simplify) + end). + Lemma transl_inop_correct: forall (s : list RTL.stackframe) (f : RTL.function) (sp : Values.val) (pc : positive) (rs : RTL.regset) (m : mem) (pc' : RTL.node), @@ -445,46 +492,8 @@ Section CORRECTNESS. econstructor. econstructor. econstructor. - simplify. - - unfold Verilog.merge_regs. - unfold_merge. apply AssocMap.gss. - - (* prove match_state *) - rewrite assumption_32bit. - econstructor; simplify; eauto. - - unfold Verilog.merge_regs. - unfold_merge. simpl. apply regs_lessdef_add_greater. apply greater_than_max_func. - assumption. - unfold Verilog.merge_regs. - unfold state_st_wf. inversion 1. subst. unfold_merge. apply AssocMap.gss. - - (* prove match_arrs *) - invert MARR. simplify. - unfold HTL.empty_stack. simplify. unfold Verilog.merge_arrs. - econstructor. - simplify. repeat split. - rewrite AssocMap.gcombine. - 2: { reflexivity. } - rewrite AssocMap.gss. - unfold Verilog.merge_arr. - setoid_rewrite H5. - reflexivity. - - rewrite combine_length. - unfold arr_repeat. simplify. - rewrite list_repeat_len. - reflexivity. - - unfold arr_repeat. simplify. - rewrite list_repeat_len; auto. - intros. - erewrite array_get_error_equal. - eauto. apply combine_none. - - assumption. + all: invert MARR; big_tac. Unshelve. constructor. @@ -598,6 +607,12 @@ Section CORRECTNESS. | [ _ : context[match ?x with | _ => _ end] |- _ ] => destruct x end. + Ltac inv_arr_access := + match goal with + | [ _ : translate_arr_access ?chunk ?addr ?args _ _ = OK ?c _ _ |- _] => + destruct c, chunk, addr, args; simplify; tac; simplify + end. + Lemma transl_iload_correct: forall (s : list RTL.stackframe) (f : RTL.function) (sp : Values.val) (pc : positive) (rs : Registers.Regmap.t Values.val) (m : mem) (chunk : AST.memory_chunk) @@ -613,9 +628,7 @@ Section CORRECTNESS. match_states (RTL.State s f sp pc' (Registers.Regmap.set dst v rs) m) R2. Proof. intros s f sp pc rs m chunk addr args dst pc' a v H H0 H1 R1 MSTATE. - inv_state. - - destruct c, chunk, addr, args; simplify; tac; simplify. + inv_state. inv_arr_access. + (** Preamble *) invert MARR. simplify. @@ -656,8 +669,7 @@ Section CORRECTNESS. rewrite Integers.Ptrofs.signed_repr; try assumption. admit. (* FIXME: Register bounds. *) apply PtrofsExtra.of_int_mod. - rewrite Integers.Int.signed_repr; simplify; try split; try assumption. - } + rewrite Integers.Int.signed_repr; simplify; try split; try assumption. } (** Read bounds proof *) assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH. @@ -693,13 +705,13 @@ Section CORRECTNESS. rewrite Integers.Ptrofs.unsigned_repr; simplify. apply Zmult_lt_reg_r with (p := 4); try lia. repeat rewrite ZLib.div_mul_undo; try lia. - split. apply Z.div_pos; try lia; apply Integers.Ptrofs.unsigned_range_2. apply Z.div_le_upper_bound; lia. } inversion NORMALISE_BOUND as [ NORMALISE_BOUND_LOW NORMALISE_BOUND_HIGH ]; clear NORMALISE_BOUND. + (** Start of proof proper *) eexists. split. eapply Smallstep.plus_one. eapply HTL.step_module; eauto. @@ -711,27 +723,10 @@ Section CORRECTNESS. econstructor. econstructor. econstructor. econstructor. econstructor. econstructor. econstructor. econstructor. - all: simplify. - - (** Verilog array lookup *) - unfold Verilog.arr_assocmap_lookup. setoid_rewrite H5. - f_equal. - - (** State Lookup *) - unfold Verilog.merge_regs. - simplify. - unfold_merge. - rewrite AssocMap.gso. - apply AssocMap.gss. - apply st_greater_than_res. - - (** Match states *) - rewrite assumption_32bit. - econstructor; eauto. + all: big_tac. (** Match assocmaps *) - unfold Verilog.merge_regs. simplify. unfold_merge. - apply regs_lessdef_add_match. + apply regs_lessdef_add_match; big_tac. (** Equality proof *) assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity. @@ -762,51 +757,7 @@ Section CORRECTNESS. rewrite H1 in I. invert I. assumption. - (** PC match *) - apply regs_lessdef_add_greater. - apply greater_than_max_func. - assumption. - - (** States well formed *) - unfold state_st_wf. inversion 1. simplify. - unfold Verilog.merge_regs. - unfold_merge. rewrite AssocMap.gso. - apply AssocMap.gss. - apply st_greater_than_res. - - (** Match arrays *) - econstructor. - repeat split; simplify. - unfold HTL.empty_stack. - simplify. - unfold Verilog.merge_arrs. - - rewrite AssocMap.gcombine. - 2: { reflexivity. } - rewrite AssocMap.gss. - unfold Verilog.merge_arr. - setoid_rewrite H5. - reflexivity. - - rewrite combine_length. - unfold arr_repeat. simplify. - rewrite list_repeat_len. - reflexivity. - - unfold arr_repeat. simplify. - rewrite list_repeat_len. - congruence. - - intros. - erewrite array_get_error_equal. - eauto. apply combine_none. - assumption. - (** RSBP preservation *) - unfold reg_stack_based_pointers. intros. - destruct (Pos.eq_dec r1 dst); try rewrite e. (* FIXME: Prepare this for automation *) - - rewrite Registers.Regmap.gss. unfold arr_stack_based_pointers in ASBP. specialize (ASBP (Integers.Ptrofs.unsigned (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4)))). @@ -819,9 +770,6 @@ Section CORRECTNESS. rewrite Integers.Ptrofs.add_zero_l in I. rewrite H1 in I. assumption. - simplify. - - rewrite Registers.Regmap.gso; auto. + (** Preamble *) invert MARR. simplify. @@ -878,8 +826,7 @@ Section CORRECTNESS. exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *) admit. (* FIXME: Register bounds. *) rewrite Integers.Int.signed_repr; simplify; try split; try assumption. - rewrite Integers.Int.signed_repr; simplify; try split; try assumption. - } + rewrite Integers.Int.signed_repr; simplify; try split; try assumption. } (** Read bounds proof *) assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH. @@ -915,7 +862,6 @@ Section CORRECTNESS. rewrite Integers.Ptrofs.unsigned_repr; simplify. apply Zmult_lt_reg_r with (p := 4); try lia. repeat rewrite ZLib.div_mul_undo; try lia. - split. apply Z.div_pos; try lia; apply Integers.Ptrofs.unsigned_range_2. apply Z.div_le_upper_bound; lia. } @@ -939,27 +885,10 @@ Section CORRECTNESS. econstructor. econstructor. econstructor. econstructor. econstructor. econstructor. - all: simplify. - - (** Verilog array lookup *) - unfold Verilog.arr_assocmap_lookup. setoid_rewrite H5. - f_equal. - - (** State Lookup *) - unfold Verilog.merge_regs. - simplify. - unfold_merge. - rewrite AssocMap.gso. - apply AssocMap.gss. - apply st_greater_than_res. - - (** Match states *) - rewrite assumption_32bit. - econstructor; eauto. + all: big_tac. (** Match assocmaps *) - unfold Verilog.merge_regs. simplify. unfold_merge. - apply regs_lessdef_add_match. + apply regs_lessdef_add_match; big_tac. (** Equality proof *) assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity. @@ -990,51 +919,7 @@ Section CORRECTNESS. rewrite H1 in I. invert I. assumption. - (** PC match *) - apply regs_lessdef_add_greater. - apply greater_than_max_func. - assumption. - - (** States well formed *) - unfold state_st_wf. inversion 1. simplify. - unfold Verilog.merge_regs. - unfold_merge. rewrite AssocMap.gso. - apply AssocMap.gss. - apply st_greater_than_res. - - (** Match arrays *) - econstructor. - repeat split; simplify. - unfold HTL.empty_stack. - simplify. - unfold Verilog.merge_arrs. - - rewrite AssocMap.gcombine. - 2: { reflexivity. } - rewrite AssocMap.gss. - unfold Verilog.merge_arr. - setoid_rewrite H5. - reflexivity. - - rewrite combine_length. - unfold arr_repeat. simplify. - rewrite list_repeat_len. - reflexivity. - - unfold arr_repeat. simplify. - rewrite list_repeat_len. - congruence. - - intros. - erewrite array_get_error_equal. - eauto. apply combine_none. - assumption. - (** RSBP preservation *) - unfold reg_stack_based_pointers. intros. - destruct (Pos.eq_dec r2 dst); try rewrite e. (* FIXME: Prepare this for automation *) - - rewrite Registers.Regmap.gss. unfold arr_stack_based_pointers in ASBP. specialize (ASBP (Integers.Ptrofs.unsigned (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4)))). @@ -1047,9 +932,6 @@ Section CORRECTNESS. rewrite Integers.Ptrofs.add_zero_l in I. rewrite H1 in I. assumption. - simplify. - - rewrite Registers.Regmap.gso; auto. + invert MARR. simplify. @@ -1102,7 +984,6 @@ Section CORRECTNESS. rewrite Integers.Ptrofs.unsigned_repr; simplify. apply Zmult_lt_reg_r with (p := 4); try lia. repeat rewrite ZLib.div_mul_undo; try lia. - split. apply Z.div_pos; try lia; apply Integers.Ptrofs.unsigned_range_2. apply Z.div_le_upper_bound; lia. } @@ -1117,27 +998,10 @@ Section CORRECTNESS. econstructor. econstructor. econstructor. simplify. econstructor. econstructor. econstructor. econstructor. simplify. - all: simplify. - - (** Verilog array lookup *) - unfold Verilog.arr_assocmap_lookup. setoid_rewrite H5. - f_equal. - - (** State Lookup *) - unfold Verilog.merge_regs. - simplify. - unfold_merge. - rewrite AssocMap.gso. - apply AssocMap.gss. - apply st_greater_than_res. - - (** Match states *) - rewrite assumption_32bit. - econstructor; eauto. + all: big_tac. (** Match assocmaps *) - unfold Verilog.merge_regs. simplify. unfold_merge. - apply regs_lessdef_add_match. + apply regs_lessdef_add_match; big_tac. (** Equality proof *) assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity. @@ -1167,51 +1031,7 @@ Section CORRECTNESS. rewrite H1 in I. invert I. assumption. - (** PC match *) - apply regs_lessdef_add_greater. - apply greater_than_max_func. - assumption. - - (** States well formed *) - unfold state_st_wf. inversion 1. simplify. - unfold Verilog.merge_regs. - unfold_merge. rewrite AssocMap.gso. - apply AssocMap.gss. - apply st_greater_than_res. - - (** Match arrays *) - econstructor. - repeat split; simplify. - unfold HTL.empty_stack. - simplify. - unfold Verilog.merge_arrs. - - rewrite AssocMap.gcombine. - 2: { reflexivity. } - rewrite AssocMap.gss. - unfold Verilog.merge_arr. - setoid_rewrite H5. - reflexivity. - - rewrite combine_length. - unfold arr_repeat. simplify. - rewrite list_repeat_len. - reflexivity. - - unfold arr_repeat. simplify. - rewrite list_repeat_len. - congruence. - - intros. - erewrite array_get_error_equal. - eauto. apply combine_none. - assumption. - (** RSBP preservation *) - unfold reg_stack_based_pointers. intros. - destruct (Pos.eq_dec r0 dst); try rewrite e. (* FIXME: Prepare this for automation *) - - rewrite Registers.Regmap.gss. unfold arr_stack_based_pointers in ASBP. specialize (ASBP (Integers.Ptrofs.unsigned (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4)))). @@ -1224,9 +1044,6 @@ Section CORRECTNESS. rewrite Integers.Ptrofs.add_zero_l in I. rewrite H1 in I. assumption. - simplify. - - rewrite Registers.Regmap.gso; auto. Admitted. Hint Resolve transl_iload_correct : htlproof. @@ -1244,9 +1061,8 @@ Section CORRECTNESS. Smallstep.plus HTL.step tge R1 Events.E0 R2 /\ match_states (RTL.State s f sp pc' rs m') R2. Proof. intros s f sp pc rs m chunk addr args src pc' a m' H H0 H1 R1 MSTATES. - inv_state. + inv_state. inv_arr_access. - destruct c, chunk, addr, args; simplify; tac; simplify. + (** Preamble *) invert MARR. simplify. @@ -1500,8 +1316,8 @@ Section CORRECTNESS. rewrite Integers.Ptrofs.add_zero_l. rewrite Integers.Ptrofs.unsigned_repr; simplify; try lia. apply ZExtra.mod_0_bounds; simplify; try lia. } - simplify. split. - exploit (BOUNDS ptr); try lia. intros. simplify. assumption. + simplify. + exploit (BOUNDS ptr); try lia. intros. simplify. exploit (BOUNDS ptr v); try lia. intros. invert H0. match goal with | |- ?x = _ => destruct x eqn:EQ end; try reflexivity. @@ -1792,8 +1608,8 @@ Section CORRECTNESS. rewrite Integers.Ptrofs.add_zero_l. rewrite Integers.Ptrofs.unsigned_repr; simplify; try lia. apply ZExtra.mod_0_bounds; simplify; try lia. } - simplify. split. - exploit (BOUNDS ptr); try lia. intros. simplify. assumption. + simplify. + exploit (BOUNDS ptr); try lia. intros. simplify. exploit (BOUNDS ptr v); try lia. intros. invert H0. match goal with | |- ?x = _ => destruct x eqn:EQ end; try reflexivity. @@ -2032,8 +1848,8 @@ Section CORRECTNESS. rewrite Integers.Ptrofs.add_zero_l. rewrite Integers.Ptrofs.unsigned_repr; simplify; try lia. apply ZExtra.mod_0_bounds; simplify; try lia. } - simplify. split. - exploit (BOUNDS ptr); try lia. intros. simplify. assumption. + simplify. + exploit (BOUNDS ptr); try lia. intros. simplify. exploit (BOUNDS ptr v); try lia. intros. invert H0. match goal with | |- ?x = _ => destruct x eqn:EQ end; try reflexivity. @@ -2071,7 +1887,6 @@ Section CORRECTNESS. apply assumption_32bit. eapply Verilog.stmnt_runp_Vnonblock_reg with (rhsval := if b then posToValue 32 ifso else posToValue 32 ifnot). - constructor. simpl. @@ -2085,6 +1900,7 @@ Section CORRECTNESS. constructor. apply boolToValue_ValueToBool. constructor. + unfold Verilog.merge_regs. unfold_merge. apply AssocMap.gss. @@ -2144,33 +1960,7 @@ Section CORRECTNESS. apply AssocMap.gss. (** Match arrays *) - invert MARR. simplify. - econstructor. - repeat split; simplify. - unfold HTL.empty_stack. - simplify. - unfold Verilog.merge_arrs. - - rewrite AssocMap.gcombine. - 2: { reflexivity. } - rewrite AssocMap.gss. - unfold Verilog.merge_arr. - setoid_rewrite H2. - reflexivity. - - rewrite combine_length. - unfold arr_repeat. simplify. - rewrite list_repeat_len. - reflexivity. - - unfold arr_repeat. simplify. - rewrite list_repeat_len. - congruence. - - intros. - erewrite array_get_error_equal. - eauto. apply combine_none. - assumption. + all: invert MARR. big_tac. Unshelve. constructor. @@ -2222,12 +2012,7 @@ Section CORRECTNESS. constructor. constructor. - unfold Verilog.merge_regs. - unfold_merge. simpl. - rewrite AssocMap.gso. - rewrite AssocMap.gso. - unfold state_st_wf in WF. eapply WF. reflexivity. - apply st_greater_than_res. apply st_greater_than_res. + unfold state_st_wf in WF; big_tac; eauto. apply HTL.step_finish. unfold Verilog.merge_regs. @@ -2249,18 +2034,10 @@ Section CORRECTNESS. constructor. econstructor; simpl; trivial. econstructor; simpl; trivial. + constructor. constructor. constructor. + constructor. constructor. constructor. - constructor. constructor. - - constructor. - econstructor; simpl; trivial. - apply Verilog.erun_Vvar. trivial. - unfold Verilog.merge_regs. - unfold_merge. simpl. - rewrite AssocMap.gso. - rewrite AssocMap.gso. - unfold state_st_wf in WF. eapply WF. trivial. - apply st_greater_than_res. apply st_greater_than_res. trivial. + unfold state_st_wf in WF; big_tac; eauto. apply HTL.step_finish. unfold Verilog.merge_regs. @@ -2303,18 +2080,14 @@ Section CORRECTNESS. apply match_state with (sp' := stk); eauto. + all: big_tac. + apply regs_lessdef_add_greater. apply greater_than_max_func. apply init_reg_assoc_empty. - unfold state_st_wf. - intros. inv H3. apply AssocMap.gss. constructor. - econstructor. simplify. - repeat split. unfold HTL.empty_stack. - simplify. apply AssocMap.gss. - unfold arr_repeat. simplify. apply list_repeat_len. intros. destruct (Mem.load AST.Mint32 m' stk -- cgit From 24b07d3b719072482f609954f584232534ed93eb Mon Sep 17 00:00:00 2001 From: James Pollard Date: Tue, 30 Jun 2020 21:34:49 +0100 Subject: Remove some explicit evar instantiations. --- src/translation/HTLgenproof.v | 51 ++++++++++++++++++++++++------------------- 1 file changed, 28 insertions(+), 23 deletions(-) diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index 9f62bb9..079cc1e 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -465,8 +465,10 @@ Section CORRECTNESS. destruct (Pos.eq_dec s d) as [EQ|EQ]; [> rewrite EQ | rewrite Registers.Regmap.gso; auto] + | [ |- context[Verilog.arr_assocmap_set _ _ _ _] ] => unfold Verilog.arr_assocmap_set + | [ H : _ ! _ = Some _ |- _] => try (setoid_rewrite H; simplify) - end). + end); simplify. Lemma transl_inop_correct: forall (s : list RTL.stackframe) (f : RTL.function) (sp : Values.val) (pc : positive) @@ -718,8 +720,7 @@ Section CORRECTNESS. apply assumption_32bit. econstructor. econstructor. econstructor. simplify. econstructor. econstructor. econstructor. simplify. - eapply Verilog.erun_Vbinop with (EQ := ?[EQ1]). (* FIXME: These will be shelved and cause sadness. *) - eapply Verilog.erun_Vbinop with (EQ := ?[EQ2]). + econstructor. econstructor. econstructor. econstructor. econstructor. econstructor. econstructor. econstructor. econstructor. econstructor. @@ -744,14 +745,18 @@ Section CORRECTNESS. intros I. - assert (Z.to_nat - (Integers.Ptrofs.unsigned - (Integers.Ptrofs.divu - OFFSET - (Integers.Ptrofs.repr 4))) - = - valueToNat (vdiv (vplus asr # r0 (ZToValue 32 z) ?EQ2) (ZToValue 32 4) ?EQ1)) - as EXPR_OK by admit. + match goal with + | [ |- context [valueToNat ?x] ] => + assert (Z.to_nat + (Integers.Ptrofs.unsigned + (Integers.Ptrofs.divu + OFFSET + (Integers.Ptrofs.repr 4))) + = + valueToNat x) + as EXPR_OK by admit + end. + rewrite <- EXPR_OK. rewrite NORMALISE in I. rewrite H1 in I. @@ -875,9 +880,7 @@ Section CORRECTNESS. apply assumption_32bit. econstructor. econstructor. econstructor. simplify. econstructor. econstructor. econstructor. simplify. - eapply Verilog.erun_Vbinop with (EQ := ?[EQ3]). (* FIXME: These will be shelved and cause sadness. *) - eapply Verilog.erun_Vbinop with (EQ := ?[EQ4]). - eapply Verilog.erun_Vbinop with (EQ := ?[EQ5]). + econstructor. econstructor. econstructor. econstructor. econstructor. econstructor. econstructor. econstructor. eapply Verilog.erun_Vbinop with (EQ := ?[EQ6]). @@ -905,15 +908,17 @@ Section CORRECTNESS. apply Z.div_pos; lia. intros I. - assert (Z.to_nat - (Integers.Ptrofs.unsigned - (Integers.Ptrofs.divu - OFFSET - (Integers.Ptrofs.repr 4))) - = valueToNat - (vdiv (vplus (vplus asr # r0 (ZToValue 32 z0) ?EQ5) - (vmul asr # r1 (ZToValue 32 z) ?EQ6) ?EQ4) (ZToValue 32 4) ?EQ3)) - as EXPR_OK by admit. + match goal with + | [ |- context [valueToNat ?x] ] => + assert (Z.to_nat + (Integers.Ptrofs.unsigned + (Integers.Ptrofs.divu + OFFSET + (Integers.Ptrofs.repr 4))) + = + valueToNat x) + as EXPR_OK by admit + end. rewrite <- EXPR_OK. rewrite NORMALISE in I. rewrite H1 in I. -- cgit From 0e0c64bf93f33044d299bfd5456d9a6b00992a0d Mon Sep 17 00:00:00 2001 From: James Pollard Date: Wed, 1 Jul 2020 20:09:15 +0100 Subject: Improve (?) automation. --- src/common/Coquplib.v | 19 +- src/common/IntegerExtra.v | 33 +- src/translation/HTLgenproof.v | 708 ++++++++++++++++++------------------------ src/translation/HTLgenspec.v | 2 +- src/verilog/Array.v | 69 ++-- 5 files changed, 380 insertions(+), 451 deletions(-) diff --git a/src/common/Coquplib.v b/src/common/Coquplib.v index ba0a5dc..c9361c2 100644 --- a/src/common/Coquplib.v +++ b/src/common/Coquplib.v @@ -23,12 +23,14 @@ From Coq Require Export List Bool. +Require Import Lia. + From coqup Require Import Show. (* Depend on CompCert for the basic library, as they declare and prove some useful theorems. *) From compcert.lib Require Export Coqlib. -From compcert Require Integers. +From compcert Require Import Integers. Ltac unfold_rec c := unfold c; fold c. @@ -44,18 +46,21 @@ Ltac solve_by_invert := solve_by_inverts 1. Ltac invert x := inversion x; subst; clear x. +Ltac destruct_match := + match goal with | [ |- context[match ?x with | _ => _ end ] ] => destruct x end. + Ltac clear_obvious := repeat match goal with | [ H : ex _ |- _ ] => invert H - | [ H : ?C _ = ?C _ |- _ ] => invert H + | [ H : Some _ = Some _ |- _ ] => invert H | [ H : _ /\ _ |- _ ] => invert H end. Ltac nicify_goals := repeat match goal with | [ |- _ /\ _ ] => split - | [ |- Some _ = Some _ ] => try reflexivity - | [ _ : ?x |- ?x ] => assumption + | [ |- Some _ = Some _ ] => f_equal + | [ |- S _ = S _ ] => f_equal end. Ltac kill_bools := @@ -124,9 +129,9 @@ Ltac unfold_constants := end end. -Ltac simplify := unfold_constants; simpl in *; - repeat (clear_obvious; nicify_goals; kill_bools); - simpl in *; try discriminate. +Ltac crush := intros; unfold_constants; simpl in *; + repeat (clear_obvious; nicify_goals; kill_bools); + simpl in *; try discriminate; try congruence; try lia; try assumption. Global Opaque Nat.div. Global Opaque Z.mul. diff --git a/src/common/IntegerExtra.v b/src/common/IntegerExtra.v index 7d3156b..6bac18d 100644 --- a/src/common/IntegerExtra.v +++ b/src/common/IntegerExtra.v @@ -27,7 +27,7 @@ Module PtrofsExtra. rewrite Zmod_mod | [ _ : _ |- context[(_ mod Ptrofs.modulus) mod m ] ] => rewrite <- Zmod_div_mod; - try (simplify; lia || assumption) + try (crush; lia || assumption) | [ _ : _ |- context[Ptrofs.modulus mod m] ] => rewrite Zdivide_mod with (a := Ptrofs.modulus); @@ -65,7 +65,7 @@ Module PtrofsExtra. | [ _ : _ |- context[_ mod Ptrofs.modulus mod m] ] => rewrite <- Zmod_div_mod; try lia; try assumption | [ _ : _ |- context[Ptrofs.unsigned _] ] => rewrite Ptrofs.unsigned_signed - end; try (simplify; lia); ptrofs_mod_tac m. + end; try crush; ptrofs_mod_tac m. Qed. Lemma of_int_mod : @@ -96,7 +96,7 @@ Module PtrofsExtra. | [ _ : _ |- context[_ mod Ptrofs.modulus mod m] ] => rewrite <- Zmod_div_mod; try lia; try assumption | [ _ : _ |- context[Ptrofs.unsigned _] ] => rewrite Ptrofs.unsigned_signed - end; try (simplify; lia); ptrofs_mod_tac m. + end; try(crush; lia); ptrofs_mod_tac m. Qed. Lemma add_mod : @@ -115,7 +115,7 @@ Module PtrofsExtra. | [ _ : _ |- context[_ mod Ptrofs.modulus mod m] ] => rewrite <- Zmod_div_mod; try lia; try assumption | [ _ : _ |- context[Ptrofs.unsigned _] ] => rewrite Ptrofs.unsigned_signed - end; try (simplify; lia); ptrofs_mod_tac m. + end; try (crush; lia); ptrofs_mod_tac m. Qed. Lemma mul_divu : @@ -156,7 +156,7 @@ Module PtrofsExtra. eapply Z.le_trans. exact H0. rewrite Z.mul_comm. - apply Z.le_mul_diag_r; simplify; lia. + apply Z.le_mul_diag_r; crush. Qed. Lemma mul_unsigned : @@ -184,6 +184,23 @@ Module PtrofsExtra. Qed. End PtrofsExtra. +Ltac ptrofs := + repeat match goal with + | [ |- context[Ptrofs.add (Ptrofs.zero) _] ] => setoid_rewrite Ptrofs.add_zero_l + | [ H : context[Ptrofs.add (Ptrofs.zero) _] |- _ ] => setoid_rewrite Ptrofs.add_zero_l in H + + | [ |- context[Ptrofs.repr 0] ] => replace (Ptrofs.repr 0) with Ptrofs.zero by reflexivity + | [ H : context[Ptrofs.repr 0] |- _ ] => + replace (Ptrofs.repr 0) with Ptrofs.zero in H by reflexivity + + | [ H: context[Ptrofs.unsigned (Ptrofs.repr (Ptrofs.unsigned _))] |- _ ] => + setoid_rewrite Ptrofs.unsigned_repr in H; [>| apply Ptrofs.unsigned_range_2] + | [ |- context[Ptrofs.unsigned (Ptrofs.repr (Ptrofs.unsigned _))] ] => + rewrite Ptrofs.unsigned_repr; [>| apply Ptrofs.unsigned_range_2] + + | [ |- context[0 <= Ptrofs.unsigned _] ] => apply Ptrofs.unsigned_range_2 + end. + Module IntExtra. Ltac int_mod_match m := @@ -202,7 +219,7 @@ Module IntExtra. rewrite Zmod_mod | [ _ : _ |- context[(_ mod Int.modulus) mod m ] ] => rewrite <- Zmod_div_mod; - try (simplify; lia || assumption) + try (crush; lia || assumption) | [ _ : _ |- context[Int.modulus mod m] ] => rewrite Zdivide_mod with (a := Int.modulus); @@ -242,7 +259,7 @@ Module IntExtra. | [ _ : _ |- context[_ mod Int.modulus mod m] ] => rewrite <- Zmod_div_mod; try lia; try assumption | [ _ : _ |- context[Int.unsigned _] ] => rewrite Int.unsigned_signed - end; try (simplify; lia); int_mod_tac m. + end; try (crush; lia); int_mod_tac m. Qed. Lemma add_mod : @@ -261,6 +278,6 @@ Module IntExtra. | [ _ : _ |- context[_ mod Int.modulus mod m] ] => rewrite <- Zmod_div_mod; try lia; try assumption | [ _ : _ |- context[Int.unsigned _] ] => rewrite Int.unsigned_signed - end; try (simplify; lia); int_mod_tac m. + end; try (crush; lia); int_mod_tac m. Qed. End IntExtra. diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index 079cc1e..f248e25 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -183,11 +183,11 @@ Lemma list_combine_none : length l = n -> list_combine Verilog.merge_cell (list_repeat None n) l = l. Proof. - induction n; intros; simplify. + induction n; intros; crush. - symmetry. apply length_zero_iff_nil. auto. - - destruct l; simplify. + - destruct l; crush. rewrite list_repeat_cons. - simplify. f_equal. + crush. f_equal. eauto. Qed. @@ -198,7 +198,7 @@ Lemma combine_none : Proof. intros. unfold combine. - simplify. + crush. rewrite <- (arr_wf a) in H. apply list_combine_none. @@ -211,12 +211,12 @@ Lemma list_combine_lookup_first : nth_error l1 n = Some None -> nth_error (list_combine Verilog.merge_cell l1 l2) n = nth_error l2 n. Proof. - induction l1; intros; simplify. + induction l1; intros; crush. rewrite nth_error_nil in H0. discriminate. - destruct l2 eqn:EQl2. simplify. + destruct l2 eqn:EQl2. crush. simpl in H. invert H. destruct n; simpl in *. invert H0. simpl. reflexivity. @@ -243,9 +243,9 @@ Lemma list_combine_lookup_second : nth_error l1 n = Some (Some x) -> nth_error (list_combine Verilog.merge_cell l1 l2) n = Some (Some x). Proof. - induction l1; intros; simplify; auto. + induction l1; intros; crush; auto. - destruct l2 eqn:EQl2. simplify. + destruct l2 eqn:EQl2. crush. simpl in H. invert H. destruct n; simpl in *. invert H0. simpl. reflexivity. @@ -432,43 +432,54 @@ Section CORRECTNESS. exists asr' asa', HTL.step tge (HTL.State res m st asr asa) Events.E0 (HTL.State res m st asr' asa'). - Ltac big_tac := - repeat (simplify; - match goal with - | [ |- context[Verilog.merge_regs _ _] ] => - unfold Verilog.merge_regs; simplify; unfold_merge - | [ |- context[_ # ?d <- _ ! ?d] ] => apply AssocMap.gss - | [ |- context[_ # ?d <- _ ! ?s] ] => rewrite AssocMap.gso; try apply st_greater_than_res - | [ |- context[valueToPos (posToValue 32 _)] ] => rewrite assumption_32bit - | [ |- context[match_states _ _] ] => econstructor; eauto - | [ |- context[Verilog.merge_arr] ] => unfold Verilog.merge_arr; simplify - | [ |- context[(AssocMap.empty _) ! _] ] => rewrite AssocMap.gempty; simplify + Opaque combine. - | [ H : ?asa ! ?r = Some _ |- Verilog.arr_assocmap_lookup ?asa ?r _ = Some _ ] => - unfold Verilog.arr_assocmap_lookup; setoid_rewrite H; f_equal - - | [ |- match_assocmaps _ _ _ # _ <- (posToValue 32 _) ] => - apply regs_lessdef_add_greater; [> apply greater_than_max_func | assumption] - - | [ |- state_st_wf _ _ ] => unfold state_st_wf; inversion 1; simplify - - | [ |- match_arrs _ _ _ _ _ ] => econstructor; simplify - | [ |- context[HTL.empty_stack] ] => unfold HTL.empty_stack; simplify - | [ |- context[Verilog.merge_arrs _ _] ] => unfold Verilog.merge_arrs; simplify - | [ |- context[(AssocMap.combine _ _ _) ! _] ] => - try (rewrite AssocMap.gcombine; [> | reflexivity]) - - | [ |- context[reg_stack_based_pointers] ] => unfold reg_stack_based_pointers; intros - | [ |- context[Registers.Regmap.get ?d (Registers.Regmap.set ?d _ _)] ] => - rewrite Registers.Regmap.gss - | [ |- context[Registers.Regmap.get ?s (Registers.Regmap.set ?d _ _)] ] => - destruct (Pos.eq_dec s d) as [EQ|EQ]; - [> rewrite EQ | rewrite Registers.Regmap.gso; auto] - - | [ |- context[Verilog.arr_assocmap_set _ _ _ _] ] => unfold Verilog.arr_assocmap_set + Ltac tac0 := + match goal with + | [ |- context[Verilog.merge_regs _ _] ] => + unfold Verilog.merge_regs; crush; unfold_merge + | [ |- context[_ # ?d <- _ ! ?d] ] => rewrite AssocMap.gss + | [ |- context[_ # ?d <- _ ! ?s] ] => rewrite AssocMap.gso + | [ |- context[valueToPos (posToValue 32 _)] ] => rewrite assumption_32bit + | [ |- context[match_states _ _] ] => econstructor; auto + | [ |- context[Verilog.merge_arr] ] => unfold Verilog.merge_arr + | [ |- context[(AssocMap.empty _) ! _] ] => rewrite AssocMap.gempty + + | [ H : ?asa ! ?r = Some _ |- Verilog.arr_assocmap_lookup ?asa ?r _ = Some _ ] => + unfold Verilog.arr_assocmap_lookup; setoid_rewrite H; f_equal + + | [ |- match_assocmaps _ _ _ # _ <- (posToValue 32 _) ] => + apply regs_lessdef_add_greater; [> apply greater_than_max_func | assumption] + + | [ |- state_st_wf _ _ ] => unfold state_st_wf; inversion 1 + + | [ |- match_arrs _ _ _ _ _ ] => econstructor; auto + | [ |- context[HTL.empty_stack] ] => unfold HTL.empty_stack + | [ |- context[Verilog.merge_arrs _ _] ] => unfold Verilog.merge_arrs + | [ |- context[(AssocMap.combine _ _ _) ! _] ] => + try (rewrite AssocMap.gcombine; [> | reflexivity]) + + | [ |- context[reg_stack_based_pointers] ] => unfold reg_stack_based_pointers; intros + | [ |- context[Registers.Regmap.get ?d (Registers.Regmap.set ?d _ _)] ] => + rewrite Registers.Regmap.gss + | [ |- context[Registers.Regmap.get ?s (Registers.Regmap.set ?d _ _)] ] => + destruct (Pos.eq_dec s d) as [EQ|EQ]; + [> rewrite EQ | rewrite Registers.Regmap.gso; auto] + + | [ |- context[Verilog.arr_assocmap_set _ _ _ _] ] => unfold Verilog.arr_assocmap_set + + | [ |- context[array_get_error _ (combine Verilog.merge_cell + (arr_repeat None _) _)] ] => + rewrite combine_lookup_first + + | [ H : opt_val_value_lessdef _ _ |- _ ] => invert H + + | [ H : context[Z.of_nat (Z.to_nat _)] |- _ ] => rewrite Z2Nat.id in H; [> solve crush |] + | [ H : _ ! _ = Some _ |- _] => setoid_rewrite H + end. - | [ H : _ ! _ = Some _ |- _] => try (setoid_rewrite H; simplify) - end); simplify. + Ltac small_tac := repeat (crush; try array; try ptrofs); crush; auto. + Ltac big_tac := repeat (crush; try array; try ptrofs; try tac0); crush; auto. Lemma transl_inop_correct: forall (s : list RTL.stackframe) (f : RTL.function) (sp : Values.val) (pc : positive) @@ -490,13 +501,12 @@ Section CORRECTNESS. apply assumption_32bit. (* processing of state *) econstructor. - simplify. + crush. econstructor. econstructor. econstructor. all: invert MARR; big_tac. - Unshelve. constructor. Qed. @@ -612,7 +622,7 @@ Section CORRECTNESS. Ltac inv_arr_access := match goal with | [ _ : translate_arr_access ?chunk ?addr ?args _ _ = OK ?c _ _ |- _] => - destruct c, chunk, addr, args; simplify; tac; simplify + destruct c, chunk, addr, args; crush; tac; crush end. Lemma transl_iload_correct: @@ -633,24 +643,24 @@ Section CORRECTNESS. inv_state. inv_arr_access. + (** Preamble *) - invert MARR. simplify. + invert MARR. crush. unfold Op.eval_addressing in H0. - destruct (Archi.ptr64) eqn:ARCHI; simplify. + destruct (Archi.ptr64) eqn:ARCHI; crush. unfold reg_stack_based_pointers in RSBP. pose proof (RSBP r0) as RSBPr0. - destruct (Registers.Regmap.get r0 rs) eqn:EQr0; simplify. + destruct (Registers.Regmap.get r0 rs) eqn:EQr0; crush. - rewrite ARCHI in H1. simplify. + rewrite ARCHI in H1. crush. subst. pose proof MASSOC as MASSOC'. invert MASSOC'. pose proof (H0 r0). assert (HPler0 : Ple r0 (RTL.max_reg_function f)) - by (eapply RTL.max_reg_function_use; eauto; simplify; eauto). + by (eapply RTL.max_reg_function_use; eauto; crush; eauto). apply H6 in HPler0. invert HPler0; try congruence. rewrite EQr0 in H8. @@ -658,7 +668,7 @@ Section CORRECTNESS. clear H0. clear H6. unfold check_address_parameter_signed in *; - unfold check_address_parameter_unsigned in *; simplify. + unfold check_address_parameter_unsigned in *; crush. remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0)) (Integers.Ptrofs.of_int (Integers.Int.repr z))) as OFFSET. @@ -666,24 +676,20 @@ Section CORRECTNESS. (** Modular preservation proof *) assert (Integers.Ptrofs.unsigned OFFSET mod 4 = 0) as MOD_PRESERVE. { rewrite HeqOFFSET. - apply PtrofsExtra.add_mod; simplify; try lia. + apply PtrofsExtra.add_mod; crush; try lia. exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *) rewrite Integers.Ptrofs.signed_repr; try assumption. admit. (* FIXME: Register bounds. *) apply PtrofsExtra.of_int_mod. - rewrite Integers.Int.signed_repr; simplify; try split; try assumption. } + rewrite Integers.Int.signed_repr; crush. } (** Read bounds proof *) assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH. - { destruct (Integers.Ptrofs.unsigned OFFSET assert (Z.to_nat @@ -756,40 +750,35 @@ Section CORRECTNESS. valueToNat x) as EXPR_OK by admit end. - rewrite <- EXPR_OK. - rewrite NORMALISE in I. - rewrite H1 in I. - invert I. assumption. + + specialize (H7 (Integers.Ptrofs.unsigned + (Integers.Ptrofs.divu + OFFSET + (Integers.Ptrofs.repr 4)))). + exploit H7; big_tac. (** RSBP preservation *) unfold arr_stack_based_pointers in ASBP. specialize (ASBP (Integers.Ptrofs.unsigned (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4)))). - exploit ASBP; auto; intros I. - - rewrite NORMALISE in I. - assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity. - rewrite ZERO in I. clear ZERO. - simplify. - rewrite Integers.Ptrofs.add_zero_l in I. - rewrite H1 in I. - assumption. + exploit ASBP; big_tac. + rewrite NORMALISE in H0. rewrite H1 in H0. assumption. + (** Preamble *) - invert MARR. simplify. + invert MARR. crush. unfold Op.eval_addressing in H0. - destruct (Archi.ptr64) eqn:ARCHI; simplify. + destruct (Archi.ptr64) eqn:ARCHI; crush. unfold reg_stack_based_pointers in RSBP. pose proof (RSBP r0) as RSBPr0. pose proof (RSBP r1) as RSBPr1. destruct (Registers.Regmap.get r0 rs) eqn:EQr0; - destruct (Registers.Regmap.get r1 rs) eqn:EQr1; simplify. + destruct (Registers.Regmap.get r1 rs) eqn:EQr1; crush. - rewrite ARCHI in H1. simplify. + rewrite ARCHI in H1. crush. subst. clear RSBPr1. @@ -798,7 +787,7 @@ Section CORRECTNESS. pose proof (H0 r0). pose proof (H0 r1). assert (HPler0 : Ple r0 (RTL.max_reg_function f)) - by (eapply RTL.max_reg_function_use; eauto; simplify; eauto). + by (eapply RTL.max_reg_function_use; eauto; crush; eauto). assert (HPler1 : Ple r1 (RTL.max_reg_function f)) by (eapply RTL.max_reg_function_use; eauto; simpl; auto). apply H6 in HPler0. @@ -810,7 +799,7 @@ Section CORRECTNESS. clear H0. clear H6. clear H8. unfold check_address_parameter_signed in *; - unfold check_address_parameter_unsigned in *; simplify. + unfold check_address_parameter_unsigned in *; crush. remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0)) (Integers.Ptrofs.of_int @@ -820,30 +809,26 @@ Section CORRECTNESS. (** Modular preservation proof *) assert (Integers.Ptrofs.unsigned OFFSET mod 4 = 0) as MOD_PRESERVE. { rewrite HeqOFFSET. - apply PtrofsExtra.add_mod; simplify; try lia. + apply PtrofsExtra.add_mod; crush; try lia. exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *) rewrite Integers.Ptrofs.signed_repr; try assumption. admit. (* FIXME: Register bounds. *) apply PtrofsExtra.of_int_mod. - apply IntExtra.add_mod; simplify; try lia. + apply IntExtra.add_mod; crush; try lia. exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *) - apply IntExtra.mul_mod; simplify; try lia. + apply IntExtra.mul_mod; crush; try lia. exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *) admit. (* FIXME: Register bounds. *) - rewrite Integers.Int.signed_repr; simplify; try split; try assumption. - rewrite Integers.Int.signed_repr; simplify; try split; try assumption. } + rewrite Integers.Int.signed_repr; crush. + rewrite Integers.Int.signed_repr; crush. } (** Read bounds proof *) assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH. - { destruct (Integers.Ptrofs.unsigned OFFSET assert (Z.to_nat @@ -920,32 +894,28 @@ Section CORRECTNESS. as EXPR_OK by admit end. rewrite <- EXPR_OK. - rewrite NORMALISE in I. - rewrite H1 in I. - invert I. assumption. + + specialize (H7 (Integers.Ptrofs.unsigned + (Integers.Ptrofs.divu + OFFSET + (Integers.Ptrofs.repr 4)))). + exploit H7; big_tac. (** RSBP preservation *) unfold arr_stack_based_pointers in ASBP. specialize (ASBP (Integers.Ptrofs.unsigned (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4)))). - exploit ASBP; auto; intros I. - - rewrite NORMALISE in I. - assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity. - rewrite ZERO in I. clear ZERO. - simplify. - rewrite Integers.Ptrofs.add_zero_l in I. - rewrite H1 in I. - assumption. + exploit ASBP; big_tac. + rewrite NORMALISE in H0. rewrite H1 in H0. assumption. - + invert MARR. simplify. + + invert MARR. crush. unfold Op.eval_addressing in H0. - destruct (Archi.ptr64) eqn:ARCHI; simplify. - rewrite ARCHI in H0. simplify. + destruct (Archi.ptr64) eqn:ARCHI; crush. + rewrite ARCHI in H0. crush. unfold check_address_parameter_unsigned in *; - unfold check_address_parameter_signed in *; simplify. + unfold check_address_parameter_signed in *; crush. assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity. rewrite ZERO in H1. clear ZERO. @@ -958,14 +928,9 @@ Section CORRECTNESS. (** Read bounds proof *) assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH. - { destruct (Integers.Ptrofs.unsigned OFFSET + assert (Z.to_nat + (Integers.Ptrofs.unsigned + (Integers.Ptrofs.divu + OFFSET + (Integers.Ptrofs.repr 4))) + = + valueToNat x) + as EXPR_OK by admit + end. + rewrite <- EXPR_OK. specialize (H7 (Integers.Ptrofs.unsigned (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4)))). - - exploit H7. - rewrite Z2Nat.id; eauto. - apply Z.div_pos; lia. - - intros I. - assert (Z.to_nat - (Integers.Ptrofs.unsigned - (Integers.Ptrofs.divu - OFFSET - (Integers.Ptrofs.repr 4))) - = - valueToNat (ZToValue 32 (Integers.Ptrofs.unsigned OFFSET / 4))) - as EXPR_OK by admit. - rewrite <- EXPR_OK. - rewrite NORMALISE in I. - rewrite H1 in I. - invert I. assumption. + exploit H7; big_tac. (** RSBP preservation *) unfold arr_stack_based_pointers in ASBP. specialize (ASBP (Integers.Ptrofs.unsigned (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4)))). - exploit ASBP; auto; intros I. - - rewrite NORMALISE in I. - assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity. - rewrite ZERO in I. clear ZERO. - simplify. - rewrite Integers.Ptrofs.add_zero_l in I. - rewrite H1 in I. - assumption. + exploit ASBP; big_tac. + rewrite NORMALISE in H0. rewrite H1 in H0. assumption. Admitted. Hint Resolve transl_iload_correct : htlproof. @@ -1069,24 +1022,24 @@ Section CORRECTNESS. inv_state. inv_arr_access. + (** Preamble *) - invert MARR. simplify. + invert MARR. crush. unfold Op.eval_addressing in H0. - destruct (Archi.ptr64) eqn:ARCHI; simplify. + destruct (Archi.ptr64) eqn:ARCHI; crush. unfold reg_stack_based_pointers in RSBP. pose proof (RSBP r0) as RSBPr0. - destruct (Registers.Regmap.get r0 rs) eqn:EQr0; simplify. + destruct (Registers.Regmap.get r0 rs) eqn:EQr0; crush. - rewrite ARCHI in H1. simplify. + rewrite ARCHI in H1. crush. subst. pose proof MASSOC as MASSOC'. invert MASSOC'. pose proof (H0 r0). assert (HPler0 : Ple r0 (RTL.max_reg_function f)) - by (eapply RTL.max_reg_function_use; eauto; simplify; eauto). + by (eapply RTL.max_reg_function_use; eauto; crush; eauto). apply H6 in HPler0. invert HPler0; try congruence. rewrite EQr0 in H8. @@ -1094,7 +1047,7 @@ Section CORRECTNESS. clear H0. clear H6. unfold check_address_parameter_unsigned in *; - unfold check_address_parameter_signed in *; simplify. + unfold check_address_parameter_signed in *; crush. remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0)) (Integers.Ptrofs.of_int (Integers.Int.repr z))) as OFFSET. @@ -1102,24 +1055,18 @@ Section CORRECTNESS. (** Modular preservation proof *) assert (Integers.Ptrofs.unsigned OFFSET mod 4 = 0) as MOD_PRESERVE. { rewrite HeqOFFSET. - apply PtrofsExtra.add_mod; simplify; try lia. + apply PtrofsExtra.add_mod; crush; try lia. exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *) rewrite Integers.Ptrofs.signed_repr; try assumption. admit. (* FIXME: Register bounds. *) apply PtrofsExtra.of_int_mod. - rewrite Integers.Int.signed_repr; simplify; try split; try assumption. - } + rewrite Integers.Int.signed_repr; crush. } (** Write bounds proof *) assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as WRITE_BOUND_HIGH. - { destruct (Integers.Ptrofs.unsigned OFFSET + assert (Z.to_nat + (Integers.Ptrofs.unsigned + (Integers.Ptrofs.divu + OFFSET + (Integers.Ptrofs.repr 4))) + = + valueToNat x) + as EXPR_OK by admit + end. assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity. inversion MASSOC; revert HeqOFFSET; subst; clear MASSOC; intros HeqOFFSET. econstructor. - repeat split; simplify. + repeat split; crush. unfold HTL.empty_stack. - simplify. + crush. unfold Verilog.merge_arrs. rewrite AssocMap.gcombine. @@ -1190,15 +1140,17 @@ Section CORRECTNESS. rewrite combine_length. rewrite <- array_set_len. - unfold arr_repeat. simplify. + unfold arr_repeat. crush. apply list_repeat_len. rewrite <- array_set_len. - unfold arr_repeat. simplify. + unfold arr_repeat. crush. rewrite list_repeat_len. rewrite H4. reflexivity. - intros. + remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0)) + (Integers.Ptrofs.of_int (Integers.Int.repr z))) as OFFSET. + destruct (4 * ptr ==Z Integers.Ptrofs.unsigned OFFSET). erewrite Mem.load_store_same. @@ -1213,19 +1165,19 @@ Section CORRECTNESS. simpl. assert (Ple src (RTL.max_reg_function f)) by (eapply RTL.max_reg_function_use; eauto; simpl; auto); - apply H0 in H14. - destruct (Registers.Regmap.get src rs) eqn:EQ_SRC; constructor; invert H14; eauto. + apply H0 in H13. + destruct (Registers.Regmap.get src rs) eqn:EQ_SRC; constructor; invert H13; eauto. rewrite <- array_set_len. - unfold arr_repeat. simplify. + unfold arr_repeat. crush. rewrite list_repeat_len. auto. assert (4 * ptr / 4 = Integers.Ptrofs.unsigned OFFSET / 4) by (f_equal; assumption). - rewrite Z.mul_comm in H14. - rewrite Z_div_mult in H14; try lia. - replace 4 with (Integers.Ptrofs.unsigned (Integers.Ptrofs.repr 4)) in H14 by reflexivity. - rewrite <- PtrofsExtra.divu_unsigned in H14; unfold_constants; try lia. - rewrite H14. rewrite EXPR_OK. + rewrite Z.mul_comm in H13. + rewrite Z_div_mult in H13; try lia. + replace 4 with (Integers.Ptrofs.unsigned (Integers.Ptrofs.repr 4)) in H13 by reflexivity. + rewrite <- PtrofsExtra.divu_unsigned in H13; unfold_constants; try lia. + rewrite H13. rewrite EXPR_OK. rewrite array_get_error_set_bound. reflexivity. unfold arr_length, arr_repeat. simpl. @@ -1242,12 +1194,11 @@ Section CORRECTNESS. right. apply ZExtra.mod_0_bounds; try lia. apply ZLib.Z_mod_mult'. - invert H13. rewrite Z2Nat.id in H19; try lia. apply Zmult_lt_compat_r with (p := 4) in H19; try lia. rewrite ZLib.div_mul_undo in H19; try lia. split; try lia. - apply Z.le_trans with (m := RTL.fn_stacksize f); simplify; lia. + apply Z.le_trans with (m := RTL.fn_stacksize f); crush; lia. } rewrite <- EXPR_OK. @@ -1259,26 +1210,26 @@ Section CORRECTNESS. eapply H7; eauto. rewrite <- array_set_len. - unfold arr_repeat. simplify. + unfold arr_repeat. crush. rewrite list_repeat_len. auto. rewrite array_gso. unfold array_get_error. unfold arr_repeat. - simplify. + crush. apply list_repeat_lookup. lia. unfold_constants. intro. - apply Z2Nat.inj_iff in H14; try lia. + apply Z2Nat.inj_iff in H13; try lia. apply Z.div_pos; try lia. - apply Integers.Ptrofs.unsigned_range_2. + apply Integers.Ptrofs.unsigned_range. assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity. unfold arr_stack_based_pointers. intros. destruct (4 * ptr ==Z Integers.Ptrofs.unsigned OFFSET). - simplify. + crush. erewrite Mem.load_store_same. 2: { rewrite ZERO. rewrite Integers.Ptrofs.add_zero_l. @@ -1286,10 +1237,11 @@ Section CORRECTNESS. rewrite Integers.Ptrofs.unsigned_repr. exact H1. apply Integers.Ptrofs.unsigned_range_2. } - simplify. + crush. destruct (Registers.Regmap.get src rs) eqn:EQ_SRC; try constructor. destruct (Archi.ptr64); try discriminate. - pose proof (RSBP src). rewrite EQ_SRC in H0. assumption. + pose proof (RSBP src). rewrite EQ_SRC in H0. + assumption. simpl. erewrite Mem.load_store_other with (m1 := m). @@ -1307,7 +1259,7 @@ Section CORRECTNESS. apply Zmult_lt_compat_r with (p := 4) in H14; try lia. rewrite ZLib.div_mul_undo in H14; try lia. split; try lia. - apply Z.le_trans with (m := RTL.fn_stacksize f); simplify; lia. + apply Z.le_trans with (m := RTL.fn_stacksize f); crush; lia. } apply ASBP; assumption. @@ -1319,10 +1271,10 @@ Section CORRECTNESS. 2: { right. right. simpl. rewrite ZERO. rewrite Integers.Ptrofs.add_zero_l. - rewrite Integers.Ptrofs.unsigned_repr; simplify; try lia. - apply ZExtra.mod_0_bounds; simplify; try lia. } - simplify. - exploit (BOUNDS ptr); try lia. intros. simplify. + rewrite Integers.Ptrofs.unsigned_repr; crush; try lia. + apply ZExtra.mod_0_bounds; crush; try lia. } + crush. + exploit (BOUNDS ptr); try lia. intros. crush. exploit (BOUNDS ptr v); try lia. intros. invert H0. match goal with | |- ?x = _ => destruct x eqn:EQ end; try reflexivity. @@ -1339,19 +1291,19 @@ Section CORRECTNESS. apply X in H0. invert H0. congruence. + (** Preamble *) - invert MARR. simplify. + invert MARR. crush. unfold Op.eval_addressing in H0. - destruct (Archi.ptr64) eqn:ARCHI; simplify. + destruct (Archi.ptr64) eqn:ARCHI; crush. unfold reg_stack_based_pointers in RSBP. pose proof (RSBP r0) as RSBPr0. pose proof (RSBP r1) as RSBPr1. destruct (Registers.Regmap.get r0 rs) eqn:EQr0; - destruct (Registers.Regmap.get r1 rs) eqn:EQr1; simplify. + destruct (Registers.Regmap.get r1 rs) eqn:EQr1; crush. - rewrite ARCHI in H1. simplify. + rewrite ARCHI in H1. crush. subst. clear RSBPr1. @@ -1360,7 +1312,7 @@ Section CORRECTNESS. pose proof (H0 r0). pose proof (H0 r1). assert (HPler0 : Ple r0 (RTL.max_reg_function f)) - by (eapply RTL.max_reg_function_use; eauto; simplify; eauto). + by (eapply RTL.max_reg_function_use; eauto; crush; eauto). assert (HPler1 : Ple r1 (RTL.max_reg_function f)) by (eapply RTL.max_reg_function_use; eauto; simpl; auto). apply H6 in HPler0. @@ -1372,7 +1324,7 @@ Section CORRECTNESS. clear H0. clear H6. clear H8. unfold check_address_parameter_signed in *; - unfold check_address_parameter_unsigned in *; simplify. + unfold check_address_parameter_unsigned in *; crush. remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0)) (Integers.Ptrofs.of_int @@ -1382,31 +1334,27 @@ Section CORRECTNESS. (** Modular preservation proof *) assert (Integers.Ptrofs.unsigned OFFSET mod 4 = 0) as MOD_PRESERVE. { rewrite HeqOFFSET. - apply PtrofsExtra.add_mod; simplify; try lia. + apply PtrofsExtra.add_mod; crush; try lia. exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *) rewrite Integers.Ptrofs.signed_repr; try assumption. admit. (* FIXME: Register bounds. *) apply PtrofsExtra.of_int_mod. - apply IntExtra.add_mod; simplify; try lia. + apply IntExtra.add_mod; crush; try lia. exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *) - apply IntExtra.mul_mod; simplify; try lia. + apply IntExtra.mul_mod; crush; try lia. exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *) admit. (* FIXME: Register bounds. *) - rewrite Integers.Int.signed_repr; simplify; try split; try assumption. - rewrite Integers.Int.signed_repr; simplify; try split; try assumption. + rewrite Integers.Int.signed_repr; crush; try split; try assumption. + rewrite Integers.Int.signed_repr; crush; try split; try assumption. } (** Write bounds proof *) assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as WRITE_BOUND_HIGH. - { destruct (Integers.Ptrofs.unsigned OFFSET destruct x eqn:EQ end; try reflexivity. @@ -1630,14 +1580,14 @@ Section CORRECTNESS. (Integers.Ptrofs.repr ptr))) v). apply X in H0. invert H0. congruence. - + invert MARR. simplify. + + invert MARR. crush. unfold Op.eval_addressing in H0. - destruct (Archi.ptr64) eqn:ARCHI; simplify. - rewrite ARCHI in H0. simplify. + destruct (Archi.ptr64) eqn:ARCHI; crush. + rewrite ARCHI in H0. crush. unfold check_address_parameter_unsigned in *; - unfold check_address_parameter_signed in *; simplify. + unfold check_address_parameter_signed in *; crush. assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity. rewrite ZERO in H1. clear ZERO. @@ -1650,14 +1600,12 @@ Section CORRECTNESS. (** Write bounds proof *) assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as WRITE_BOUND_HIGH. - { destruct (Integers.Ptrofs.unsigned OFFSET destruct x eqn:EQ end; try reflexivity. @@ -1906,66 +1853,10 @@ Section CORRECTNESS. apply boolToValue_ValueToBool. constructor. - unfold Verilog.merge_regs. - unfold_merge. - apply AssocMap.gss. - - destruct b. - rewrite assumption_32bit. - simplify. - apply match_state with (sp' := sp'); eauto. - unfold Verilog.merge_regs. - unfold_merge. - apply regs_lessdef_add_greater. apply greater_than_max_func. - assumption. - - unfold state_st_wf. intros. - invert H3. - unfold Verilog.merge_regs. unfold_merge. - apply AssocMap.gss. - - (** Match arrays *) - invert MARR. simplify. - econstructor. - repeat split; simplify. - unfold HTL.empty_stack. - simplify. - unfold Verilog.merge_arrs. - - rewrite AssocMap.gcombine. - 2: { reflexivity. } - rewrite AssocMap.gss. - unfold Verilog.merge_arr. - setoid_rewrite H4. - reflexivity. - - rewrite combine_length. - unfold arr_repeat. simplify. - rewrite list_repeat_len. - reflexivity. - - unfold arr_repeat. simplify. - rewrite list_repeat_len. - congruence. - - intros. - erewrite array_get_error_equal. - eauto. apply combine_none. - assumption. - - rewrite assumption_32bit. - apply match_state with (sp' := sp'); eauto. - unfold Verilog.merge_regs. unfold_merge. - apply regs_lessdef_add_greater. apply greater_than_max_func. - assumption. - - unfold state_st_wf. intros. - invert H1. - unfold Verilog.merge_regs. unfold_merge. - apply AssocMap.gss. + big_tac. - (** Match arrays *) - all: invert MARR. big_tac. + invert MARR. + destruct b; rewrite assumption_32bit; big_tac. Unshelve. constructor. @@ -2018,6 +1909,8 @@ Section CORRECTNESS. constructor. constructor. unfold state_st_wf in WF; big_tac; eauto. + apply st_greater_than_res. + apply st_greater_than_res. apply HTL.step_finish. unfold Verilog.merge_regs. @@ -2043,6 +1936,8 @@ Section CORRECTNESS. constructor. constructor. constructor. unfold state_st_wf in WF; big_tac; eauto. + apply st_greater_than_res. + apply st_greater_than_res. apply HTL.step_finish. unfold Verilog.merge_regs. @@ -2081,7 +1976,7 @@ Section CORRECTNESS. inversion MSTATE; subst. inversion TF; subst. econstructor. split. apply Smallstep.plus_one. - eapply HTL.step_call. simplify. + eapply HTL.step_call. crush. apply match_state with (sp' := stk); eauto. @@ -2093,8 +1988,6 @@ Section CORRECTNESS. constructor. - apply list_repeat_len. - intros. destruct (Mem.load AST.Mint32 m' stk (Integers.Ptrofs.unsigned (Integers.Ptrofs.add Integers.Ptrofs.zero @@ -2103,17 +1996,20 @@ Section CORRECTNESS. pose proof H as ALLOC. eapply LOAD_ALLOC in ALLOC. 2: { exact LOAD. } + ptrofs. rewrite LOAD. rewrite ALLOC. repeat constructor. - constructor. + + ptrofs. rewrite LOAD. + repeat constructor. unfold reg_stack_based_pointers. intros. - unfold RTL.init_regs; simplify. + unfold RTL.init_regs; crush. destruct (RTL.fn_params f); rewrite Registers.Regmap.gi; constructor. unfold arr_stack_based_pointers. intros. - simplify. + crush. destruct (Mem.load AST.Mint32 m' stk (Integers.Ptrofs.unsigned (Integers.Ptrofs.add Integers.Ptrofs.zero @@ -2134,36 +2030,34 @@ Section CORRECTNESS. unfold Mem.alloc in H. invert H. - simplify. + crush. unfold Mem.load. intros. match goal with | |- context[if ?x then _ else _] => destruct x end; try congruence. invert v0. unfold Mem.range_perm in H3. - unfold Mem.perm in H3. simplify. + unfold Mem.perm in H3. crush. unfold Mem.perm_order' in H3. - rewrite Integers.Ptrofs.add_zero_l in H3. - rewrite Integers.Ptrofs.unsigned_repr in H3; simplify; try lia. - exploit (H3 ptr). lia. intros. + small_tac. + exploit (H3 ptr). rewrite Integers.Ptrofs.unsigned_repr; small_tac. intros. rewrite Maps.PMap.gss in H8. match goal with | H8 : context[if ?x then _ else _] |- _ => destruct x eqn:EQ end; try contradiction. - simplify. + crush. apply proj_sumbool_true in H10. lia. unfold Mem.alloc in H. invert H. - simplify. + crush. unfold Mem.store. intros. match goal with | |- context[if ?x then _ else _] => destruct x end; try congruence. invert v0. unfold Mem.range_perm in H3. - unfold Mem.perm in H3. simplify. + unfold Mem.perm in H3. crush. unfold Mem.perm_order' in H3. - rewrite Integers.Ptrofs.add_zero_l in H3. - rewrite Integers.Ptrofs.unsigned_repr in H3; simplify; try lia. - exploit (H3 ptr). lia. intros. + small_tac. + exploit (H3 ptr). rewrite Integers.Ptrofs.unsigned_repr; small_tac. intros. rewrite Maps.PMap.gss in H8. match goal with | H8 : context[if ?x then _ else _] |- _ => destruct x eqn:EQ end; try contradiction. - simplify. + crush. apply proj_sumbool_true in H10. lia. Opaque Mem.alloc. Opaque Mem.load. diff --git a/src/translation/HTLgenspec.v b/src/translation/HTLgenspec.v index a78256b..0cdecba 100644 --- a/src/translation/HTLgenspec.v +++ b/src/translation/HTLgenspec.v @@ -426,7 +426,7 @@ Proof. destruct (0 <=? RTL.fn_stacksize f) eqn:STACK_BOUND_LOW; destruct (RTL.fn_stacksize f nth_error (list_set i x l) i = Some x. Proof. - induction l; intros; destruct i; simplify; try lia; try reflexivity; firstorder. + induction l; intros; destruct i; crush; firstorder. Qed. Hint Resolve list_set_spec1 : array. @@ -37,7 +37,7 @@ Lemma list_set_spec2 {A : Type} : forall l i (x : A) d, i < length l -> nth i (list_set i x l) d = x. Proof. - induction l; intros; destruct i; simplify; try lia; try reflexivity; firstorder. + induction l; intros; destruct i; crush; firstorder. Qed. Hint Resolve list_set_spec2 : array. @@ -46,7 +46,7 @@ Lemma list_set_spec3 {A : Type} : i1 <> i2 -> nth_error (list_set i1 x l) i2 = nth_error l i2. Proof. - induction l; intros; destruct i1; destruct i2; simplify; try lia; try reflexivity; firstorder. + induction l; intros; destruct i1; destruct i2; crush; firstorder. Qed. Hint Resolve list_set_spec3 : array. @@ -56,7 +56,7 @@ Lemma array_set_wf {A : Type} : Proof. induction l; intros; destruct i; auto. - invert H; simplify; auto. + invert H; crush; auto. Qed. Definition array_set {A : Type} (i : nat) (x : A) (a : Array A) := @@ -72,7 +72,7 @@ Proof. intros. rewrite <- a.(arr_wf) in H. - unfold array_set. simplify. + unfold array_set. crush. eauto with array. Qed. Hint Resolve array_set_spec1 : array. @@ -84,7 +84,7 @@ Proof. intros. rewrite <- a.(arr_wf) in H. - unfold array_set. simplify. + unfold array_set. crush. eauto with array. Qed. Hint Resolve array_set_spec2 : array. @@ -93,7 +93,7 @@ Lemma array_set_len {A : Type} : forall a i (x : A), a.(arr_length) = (array_set i x a).(arr_length). Proof. - unfold array_set. simplify. reflexivity. + unfold array_set. crush. Qed. Definition array_get_error {A : Type} (i : nat) (a : Array A) : option A := @@ -104,7 +104,7 @@ Lemma array_get_error_equal {A : Type} : a.(arr_contents) = b.(arr_contents) -> array_get_error i a = array_get_error i b. Proof. - unfold array_get_error. congruence. + unfold array_get_error. crush. Qed. Lemma array_get_error_bound {A : Type} : @@ -142,7 +142,7 @@ Proof. unfold array_get_error. unfold array_set. - simplify. + crush. eauto with array. Qed. @@ -180,18 +180,17 @@ Fixpoint list_repeat' {A : Type} (acc : list A) (a : A) (n : nat) : list A := Lemma list_repeat'_len {A : Type} : forall (a : A) n l, length (list_repeat' l a n) = (n + Datatypes.length l)%nat. Proof. - induction n; intros; simplify; try reflexivity. + induction n; intros; crush; try reflexivity. specialize (IHn (a :: l)). rewrite IHn. - simplify. - lia. + crush. Qed. Lemma list_repeat'_app {A : Type} : forall (a : A) n l, list_repeat' l a n = list_repeat' [] a n ++ l. Proof. - induction n; intros; simplify; try reflexivity. + induction n; intros; crush; try reflexivity. pose proof IHn. specialize (H (a :: l)). @@ -207,7 +206,7 @@ Qed. Lemma list_repeat'_head_tail {A : Type} : forall n (a : A), a :: list_repeat' [] a n = list_repeat' [] a n ++ [a]. Proof. - induction n; intros; simplify; try reflexivity. + induction n; intros; crush; try reflexivity. rewrite list_repeat'_app. replace (a :: list_repeat' [] a n ++ [a]) with (list_repeat' [] a n ++ [a] ++ [a]). @@ -232,17 +231,17 @@ Proof. intros. unfold list_repeat. rewrite list_repeat'_len. - simplify. lia. + crush. Qed. Lemma dec_list_repeat_spec {A : Type} : forall n (a : A) a', (forall x x' : A, {x' = x} + {~ x' = x}) -> In a' (list_repeat a n) -> a' = a. Proof. - induction n; intros; simplify; try contradiction. + induction n; intros; crush. unfold list_repeat in *. - simplify. + crush. rewrite list_repeat'_app in H. pose proof (X a a'). @@ -278,14 +277,19 @@ Lemma list_repeat_lookup {A : Type} : Proof. induction n; intros. - destruct i; simplify; lia. + destruct i; crush. rewrite list_repeat_cons. - destruct i; simplify; firstorder. + destruct i; crush; firstorder. Qed. Definition arr_repeat {A : Type} (a : A) (n : nat) : Array A := make_array (list_repeat a n). +Lemma arr_repeat_length {A : Type} : forall n (a : A), arr_length (arr_repeat a n) = n. +Proof. + unfold list_repeat. crush. apply list_repeat_len. +Qed. + Fixpoint list_combine {A B C : Type} (f : A -> B -> C) (x : list A) (y : list B) : list C := match x, y with | a :: t, b :: t' => f a b :: list_combine f t t' @@ -295,9 +299,9 @@ Fixpoint list_combine {A B C : Type} (f : A -> B -> C) (x : list A) (y : list B) Lemma list_combine_length {A B C : Type} (f : A -> B -> C) : forall (x : list A) (y : list B), length (list_combine f x y) = min (length x) (length y). Proof. - induction x; intros; simplify; try reflexivity. + induction x; intros; crush. - destruct y; simplify; auto. + destruct y; crush; auto. Qed. Definition combine {A B C : Type} (f : A -> B -> C) (x : Array A) (y : Array B) : Array C := @@ -310,15 +314,24 @@ Proof. unfold combine. unfold make_array. - simplify. + crush. rewrite <- (arr_wf x) in *. rewrite <- (arr_wf y) in *. - destruct (arr_contents x); simplify. - - reflexivity. - - destruct (arr_contents y); simplify. - f_equal. - rewrite list_combine_length. - destruct (Min.min_dec (length l) (length l0)); congruence. + destruct (arr_contents x); destruct (arr_contents y); crush. + rewrite list_combine_length. + destruct (Min.min_dec (length l) (length l0)); congruence. Qed. + +Ltac array := + try match goal with + | [ |- context[arr_length (combine _ _ _)] ] => + rewrite combine_length + | [ |- context[length (list_repeat _ _)] ] => + rewrite list_repeat_len + | |- context[array_get_error _ (arr_repeat ?x _) = Some ?x] => + unfold array_get_error, arr_repeat + | |- context[nth_error (list_repeat ?x _) _ = Some ?x] => + apply list_repeat_lookup + end. -- cgit From 995ab555d848fcf6188734e6b46677131d4cc173 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Wed, 1 Jul 2020 20:19:04 +0100 Subject: Tidy up (?) automation slightly... --- src/translation/HTLgenproof.v | 37 +++++++++++++++++-------------------- 1 file changed, 17 insertions(+), 20 deletions(-) diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index f248e25..6e470d5 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -436,44 +436,41 @@ Section CORRECTNESS. Ltac tac0 := match goal with - | [ |- context[Verilog.merge_regs _ _] ] => - unfold Verilog.merge_regs; crush; unfold_merge - | [ |- context[_ # ?d <- _ ! ?d] ] => rewrite AssocMap.gss - | [ |- context[_ # ?d <- _ ! ?s] ] => rewrite AssocMap.gso | [ |- context[valueToPos (posToValue 32 _)] ] => rewrite assumption_32bit - | [ |- context[match_states _ _] ] => econstructor; auto + + | [ |- context[Verilog.merge_arrs _ _] ] => unfold Verilog.merge_arrs | [ |- context[Verilog.merge_arr] ] => unfold Verilog.merge_arr + | [ |- context[Verilog.merge_regs _ _] ] => unfold Verilog.merge_regs; crush; unfold_merge + | [ |- context[reg_stack_based_pointers] ] => unfold reg_stack_based_pointers; intros + | [ |- context[Verilog.arr_assocmap_set _ _ _ _] ] => unfold Verilog.arr_assocmap_set + + | [ |- context[HTL.empty_stack] ] => unfold HTL.empty_stack + + | [ |- context[_ # ?d <- _ ! ?d] ] => rewrite AssocMap.gss + | [ |- context[_ # ?d <- _ ! ?s] ] => rewrite AssocMap.gso | [ |- context[(AssocMap.empty _) ! _] ] => rewrite AssocMap.gempty - | [ H : ?asa ! ?r = Some _ |- Verilog.arr_assocmap_lookup ?asa ?r _ = Some _ ] => - unfold Verilog.arr_assocmap_lookup; setoid_rewrite H; f_equal + | [ |- context[array_get_error _ (combine Verilog.merge_cell (arr_repeat None _) _)] ] => + rewrite combine_lookup_first + | [ |- state_st_wf _ _ ] => unfold state_st_wf; inversion 1 + | [ |- context[match_states _ _] ] => econstructor; auto + | [ |- match_arrs _ _ _ _ _ ] => econstructor; auto | [ |- match_assocmaps _ _ _ # _ <- (posToValue 32 _) ] => apply regs_lessdef_add_greater; [> apply greater_than_max_func | assumption] - | [ |- state_st_wf _ _ ] => unfold state_st_wf; inversion 1 - - | [ |- match_arrs _ _ _ _ _ ] => econstructor; auto - | [ |- context[HTL.empty_stack] ] => unfold HTL.empty_stack - | [ |- context[Verilog.merge_arrs _ _] ] => unfold Verilog.merge_arrs + | [ H : ?asa ! ?r = Some _ |- Verilog.arr_assocmap_lookup ?asa ?r _ = Some _ ] => + unfold Verilog.arr_assocmap_lookup; setoid_rewrite H; f_equal | [ |- context[(AssocMap.combine _ _ _) ! _] ] => try (rewrite AssocMap.gcombine; [> | reflexivity]) - | [ |- context[reg_stack_based_pointers] ] => unfold reg_stack_based_pointers; intros | [ |- context[Registers.Regmap.get ?d (Registers.Regmap.set ?d _ _)] ] => rewrite Registers.Regmap.gss | [ |- context[Registers.Regmap.get ?s (Registers.Regmap.set ?d _ _)] ] => destruct (Pos.eq_dec s d) as [EQ|EQ]; [> rewrite EQ | rewrite Registers.Regmap.gso; auto] - | [ |- context[Verilog.arr_assocmap_set _ _ _ _] ] => unfold Verilog.arr_assocmap_set - - | [ |- context[array_get_error _ (combine Verilog.merge_cell - (arr_repeat None _) _)] ] => - rewrite combine_lookup_first - | [ H : opt_val_value_lessdef _ _ |- _ ] => invert H - | [ H : context[Z.of_nat (Z.to_nat _)] |- _ ] => rewrite Z2Nat.id in H; [> solve crush |] | [ H : _ ! _ = Some _ |- _] => setoid_rewrite H end. -- cgit