From 86e64fd05cea7b1da996701cd3653db5f471f8d1 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 10 Aug 2023 12:09:03 +0100 Subject: Finish final forward simulation correctness --- src/Compiler.v | 263 ++- src/common/IntegerExtra.v | 20 - src/hls/CondElimproof.v | 6 + src/hls/DMemorygen.v | 56 - src/hls/GiblePargenproof.v | 6 + src/hls/GiblePargenproofCommon.v | 115 -- src/hls/GiblePargenproofEquiv.v | 124 -- src/hls/GiblePargenproofEvaluable.v | 34 - src/hls/GibleSeqgenproof.v | 6 + src/hls/HTLgenproof.v | 2941 -------------------------------- src/hls/IfConversion.v | 10 +- src/hls/IfConversionproof.v | 82 +- src/hls/Memorygen.v | 3201 ----------------------------------- src/hls/Predicate.v | 61 - src/hls/Value.v | 552 ------ 15 files changed, 172 insertions(+), 7305 deletions(-) delete mode 100644 src/hls/HTLgenproof.v delete mode 100644 src/hls/Memorygen.v delete mode 100644 src/hls/Value.v (limited to 'src') diff --git a/src/Compiler.v b/src/Compiler.v index d76db43..deb7da2 100644 --- a/src/Compiler.v +++ b/src/Compiler.v @@ -62,6 +62,7 @@ Require vericert.hls.Veriloggen. Require vericert.hls.Veriloggenproof. Require vericert.hls.HTLgen. Require vericert.hls.GibleSeq. +Require vericert.hls.GiblePar. Require vericert.hls.GibleSeqgen. Require vericert.hls.GiblePargen. Require vericert.hls.HTLPargen. @@ -72,11 +73,13 @@ Require vericert.hls.CondElim. Require vericert.hls.DeadBlocks. (*Require vericert.hls.PipelineOp.*) Require vericert.HLSOpts. -Require vericert.hls.Memorygen. Require vericert.hls.DMemorygen. Require vericert.hls.ClockRegisters. - -Require Import vericert.hls.HTLgenproof. +Require vericert.hls.GibleSeqgenproof. +Require vericert.hls.CondElimproof. +Require vericert.hls.IfConversionproof. +Require vericert.hls.DeadBlocksproof. +Require vericert.hls.GiblePargenproof. (*| Declarations @@ -147,18 +150,6 @@ Definition match_if {A: Type} (flag: unit -> bool) (R: A -> A -> Prop) Definition match_rep {A: Type} (R: A -> A -> Prop): A -> A -> Prop := Relation_Operators.clos_refl_trans A R. -Global Instance TransfIfLink {A: Type} {LA: Linker A} - (transf: A -> A -> Prop) (TL: TransfLink transf) - : TransfLink (match_rep transf). -Admitted. - -Lemma total_rep_match: - forall (A B: Type) (n: list B) (f: A -> B -> A) - (rel: A -> A -> Prop) (prog: A), - (forall b p, rel p (f p b)) -> - match_rep rel prog (fold_left f n prog). -Proof. Admitted. - Lemma total_if_match: forall (A: Type) (flag: unit -> bool) (f: A -> A) (rel: A -> A -> Prop) (prog: A), @@ -201,17 +192,17 @@ Proof. Qed. (*| -Top-level Translation ---------------------- - -An optimised transformation function from ``RTL`` to ``Verilog`` can then be -defined, which applies the front end compiler optimisations of CompCert to the -RTL that is generated and then performs the two Vericert passes from RTL to HTL -and then from HTL to Verilog. +This is an unverified version of transf_hls with some experimental additions +such as scheduling that aren't completed yet. |*) - -Definition transf_backend (r : RTL.program) : res Verilog.program := - OK r +Definition transf_hls_temp (p : Csyntax.program) : res Verilog.program := + OK p + @@@ SimplExpr.transl_program + @@@ SimplLocals.transf_program + @@@ Cshmgen.transl_program + @@@ Cminorgen.transl_program + @@@ Selection.sel_program + @@@ RTLgen.transl_program @@@ Inlining.transf_program @@ print (print_RTL 1) @@ Renumber.transf_program @@ -229,33 +220,30 @@ Definition transf_backend (r : RTL.program) : res Verilog.program := @@ print (print_RTL 6) @@@ time "Unused globals" Unusedglob.transform_program @@ print (print_RTL 7) - @@@ HTLgen.transl_program - @@ print (print_HTL 0) - @@ total_if HLSOpts.optim_ram Memorygen.transf_program - @@ print (print_HTL 1) - @@ Veriloggen.transl_program. - -(*| -The transformation functions from RTL to Verilog are then added to the backend -of the CompCert transformations from Clight to RTL. -|*) - -Definition transf_hls (p : Csyntax.program) : res Verilog.program := - OK p - @@@ SimplExpr.transl_program - @@@ SimplLocals.transf_program - @@@ Cshmgen.transl_program - @@@ Cminorgen.transl_program - @@@ Selection.sel_program - @@@ RTLgen.transl_program - @@ print (print_RTL 0) - @@@ transf_backend. + @@@ GibleSeqgen.transl_program + @@ print (print_GibleSeq 0) + @@ total_if HLSOpts.optim_if_conversion CondElim.transf_program + @@ print (print_GibleSeq 1) + (* @@ total_if HLSOpts.optim_if_conversion (fold_left (fun a b => IfConversion.transf_program b a) (PTree.empty _ :: PTree.empty _ :: nil)) *) + @@ total_if HLSOpts.optim_if_conversion (IfConversion.transf_program) + @@ print (print_GibleSeq 2) + @@ total_if HLSOpts.optim_if_conversion (IfConversion.transf_program) + @@ print (print_GibleSeq 3) + @@ total_if HLSOpts.optim_if_conversion (IfConversion.transf_program) + @@ print (print_GibleSeq 4) + @@@ DeadBlocks.transf_program + @@ print (print_GibleSeq 5) + @@@ GiblePargen.transl_program + @@ print (print_GiblePar 0) + @@@ HTLPargen.transl_program + @@ print (print_DHTL 0) + @@ DMemorygen.transf_program + @@ print (print_DHTL 1) + @@@ ClockRegisters.transl_program + @@ print (print_DHTL 2) + @@ DVeriloggen.transl_program. -(*| -This is an unverified version of transf_hls with some experimental additions -such as scheduling that aren't completed yet. -|*) -Definition transf_hls_temp (p : Csyntax.program) : res Verilog.program := +Definition transf_hls_scheduled (p : Csyntax.program) : res GiblePar.GiblePar.program := OK p @@@ SimplExpr.transl_program @@@ SimplLocals.transf_program @@ -285,23 +273,16 @@ Definition transf_hls_temp (p : Csyntax.program) : res Verilog.program := @@ total_if HLSOpts.optim_if_conversion CondElim.transf_program @@ print (print_GibleSeq 1) (* @@ total_if HLSOpts.optim_if_conversion (fold_left (fun a b => IfConversion.transf_program b a) (PTree.empty _ :: PTree.empty _ :: nil)) *) - @@ total_if HLSOpts.optim_if_conversion (IfConversion.transf_program (PTree.empty _)) + @@ total_if HLSOpts.optim_if_conversion (IfConversion.transf_program) @@ print (print_GibleSeq 2) - @@ total_if HLSOpts.optim_if_conversion (IfConversion.transf_program (PTree.empty _)) + @@ total_if HLSOpts.optim_if_conversion (IfConversion.transf_program) @@ print (print_GibleSeq 3) - @@ total_if HLSOpts.optim_if_conversion (IfConversion.transf_program (PTree.empty _)) + @@ total_if HLSOpts.optim_if_conversion (IfConversion.transf_program) @@ print (print_GibleSeq 4) @@@ DeadBlocks.transf_program @@ print (print_GibleSeq 5) @@@ GiblePargen.transl_program - @@ print (print_GiblePar 0) - @@@ HTLPargen.transl_program - @@ print (print_DHTL 0) - @@ DMemorygen.transf_program - @@ print (print_DHTL 1) - @@@ ClockRegisters.transl_program - @@ print (print_DHTL 2) - @@ DVeriloggen.transl_program. + @@ print (print_GiblePar 0). (*| Correctness Proof @@ -327,10 +308,13 @@ Definition CompCert's_passes := ::: mkpass (match_if Compopts.optim_CSE CSEproof.match_prog) ::: mkpass (match_if Compopts.optim_redundancy Deadcodeproof.match_prog) ::: mkpass Unusedglobproof.match_prog - ::: (@mkpass _ _ HTLgenproof.match_prog - (HTLgenproof.TransfHTLLink HTLgen.transl_program)) - ::: mkpass (match_if HLSOpts.optim_ram Memorygen.match_prog) - ::: mkpass Veriloggenproof.match_prog + ::: mkpass GibleSeqgenproof.match_prog + ::: mkpass (match_if HLSOpts.optim_if_conversion CondElimproof.match_prog) + ::: mkpass (match_if HLSOpts.optim_if_conversion IfConversionproof.match_prog) + ::: mkpass (match_if HLSOpts.optim_if_conversion IfConversionproof.match_prog) + ::: mkpass (match_if HLSOpts.optim_if_conversion IfConversionproof.match_prog) + ::: mkpass DeadBlocksproof.match_prog + ::: mkpass GiblePargenproof.match_prog ::: pass_nil _. (*| @@ -338,7 +322,7 @@ These passes are then composed into a larger, top-level ``match_prog`` predicate which matches a C program directly with a Verilog program. |*) -Definition match_prog: Csyntax.program -> Verilog.program -> Prop := +Definition match_prog: Csyntax.program -> GiblePar.GiblePar.program -> Prop := pass_match (compose_passes CompCert's_passes). (*| @@ -348,44 +332,47 @@ is performed using the ``transf_hls`` function declared above. Theorem transf_hls_match: forall p tp, - transf_hls p = OK tp -> + transf_hls_scheduled p = OK tp -> match_prog p tp. Proof. intros p tp T. - unfold transf_hls, time in T. simpl in T. + unfold transf_hls_scheduled, time in T. cbn in T. destruct (SimplExpr.transl_program p) as [p1|e] eqn:P1; - simpl in T; try discriminate. + cbn in T; try discriminate. destruct (SimplLocals.transf_program p1) as [p2|e] eqn:P2; - simpl in T; try discriminate. + cbn in T; try discriminate. destruct (Cshmgen.transl_program p2) as [p3|e] eqn:P3; - simpl in T; try discriminate. + cbn in T; try discriminate. destruct (Cminorgen.transl_program p3) as [p4|e] eqn:P4; - simpl in T; try discriminate. + cbn in T; try discriminate. destruct (Selection.sel_program p4) as [p5|e] eqn:P5; - simpl in T; try discriminate. + cbn in T; try discriminate. rewrite ! compose_print_identity in T. destruct (RTLgen.transl_program p5) as [p6|e] eqn:P6; - simpl in T; try discriminate. - unfold transf_backend, time in T. simpl in T. - rewrite ! compose_print_identity in T. + cbn in T; try discriminate. + unfold time in T. cbn in T. destruct (Inlining.transf_program p6) as [p7|e] eqn:P7; - simpl in T; try discriminate. + cbn in T; try discriminate. set (p8 := Renumber.transf_program p7) in *. set (p9 := total_if Compopts.optim_constprop Constprop.transf_program p8) in *. set (p10 := total_if Compopts.optim_constprop Renumber.transf_program p9) in *. destruct (partial_if Compopts.optim_CSE CSE.transf_program p10) - as [p11|e] eqn:P11; simpl in T; try discriminate. + as [p11|e] eqn:P11; cbn in T; try discriminate. destruct (partial_if Compopts.optim_redundancy Deadcode.transf_program p11) - as [p12|e] eqn:P12; simpl in T; try discriminate. + as [p12|e] eqn:P12; cbn in T; try discriminate. destruct (Unusedglob.transform_program p12) as [p13|e] eqn:P13; - simpl in T; try discriminate. - destruct (HTLgen.transl_program p13) as [p14|e] eqn:P14; - simpl in T; try discriminate. - set (p15 := total_if HLSOpts.optim_ram Memorygen.transf_program p14) in *. - set (p16 := Veriloggen.transl_program p15) in *. - unfold match_prog; simpl. + cbn in T; try discriminate. + destruct (GibleSeqgen.transl_program p13) as [p14|e] eqn:P14; + cbn in T; try discriminate. + set (p15 := total_if HLSOpts.optim_if_conversion CondElim.transf_program p14) in *. + set (p16 := total_if HLSOpts.optim_if_conversion IfConversion.transf_program p15) in *. + set (p17 := total_if HLSOpts.optim_if_conversion IfConversion.transf_program p16) in *. + set (p18 := total_if HLSOpts.optim_if_conversion IfConversion.transf_program p17) in *. + destruct (DeadBlocks.transf_program p18) as [p19|e] eqn:P19; cbn in T; try discriminate. + destruct (GiblePargen.transl_program p19) as [p20|e] eqn:P20; cbn in T; try discriminate. + unfold match_prog; cbn. exists p1; split. apply SimplExprproof.transf_program_match; auto. exists p2; split. apply SimplLocalsproof.match_transf_program; auto. exists p3; split. apply Cshmgenproof.transf_program_match; auto. @@ -403,19 +390,24 @@ Proof. exists p12; split. eapply partial_if_match; eauto. apply Deadcodeproof.transf_program_match. exists p13; split. apply Unusedglobproof.transf_program_match; auto. - exists p14; split. apply HTLgenproof.transf_program_match; auto. + exists p14; split. apply GibleSeqgenproof.transf_program_match; auto. exists p15; split. apply total_if_match. - apply Memorygen.transf_program_match; auto. - exists p16; split. apply Veriloggenproof.transf_program_match; auto. + apply CondElimproof.transf_program_match; auto. + exists p16; split. apply total_if_match. + apply IfConversionproof.transf_program_match; auto. + exists p17; split. apply total_if_match. + apply IfConversionproof.transf_program_match; auto. + exists p18; split. apply total_if_match. + apply IfConversionproof.transf_program_match; auto. + exists p19; split. apply DeadBlocksproof.transf_program_match; auto. + exists p20; split. apply GiblePargenproof.transf_program_match; auto. inv T. reflexivity. Qed. Theorem cstrategy_semantic_preservation: forall p tp, match_prog p tp -> - forward_simulation (Cstrategy.semantics p) (Verilog.semantics tp) - /\ backward_simulation (atomic (Cstrategy.semantics p)) - (Verilog.semantics tp). + forward_simulation (Cstrategy.semantics p) (GiblePar.GiblePar.semantics tp). Proof. intros p tp M. unfold match_prog, pass_match in M; simpl in M. Ltac DestructM := @@ -426,7 +418,7 @@ Ltac DestructM := end. repeat DestructM. subst tp. assert (F: forward_simulation (Cstrategy.semantics p) - (Verilog.semantics p16)). + (GiblePar.GiblePar.semantics p20)). { eapply compose_forward_simulations. eapply SimplExprproof.transl_program_correct; eassumption. @@ -459,44 +451,26 @@ Ltac DestructM := eapply compose_forward_simulations. eapply Unusedglobproof.transf_program_correct; eassumption. eapply compose_forward_simulations. - eapply HTLgenproof.transf_program_correct. eassumption. + eapply GibleSeqgenproof.transf_program_correct. eassumption. eapply compose_forward_simulations. eapply match_if_simulation. eassumption. - exact Memorygen.transf_program_correct; eassumption. - eapply Veriloggenproof.transf_program_correct; eassumption. + exact CondElimproof.transf_program_correct; eassumption. + eapply compose_forward_simulations. + eapply match_if_simulation. eassumption. + exact IfConversionproof.transf_program_correct; eassumption. + eapply compose_forward_simulations. + eapply match_if_simulation. eassumption. + exact IfConversionproof.transf_program_correct; eassumption. + eapply compose_forward_simulations. + eapply match_if_simulation. eassumption. + exact IfConversionproof.transf_program_correct; eassumption. + eapply compose_forward_simulations. + eapply DeadBlocksproof.transf_program_correct; eassumption. + eapply compose_forward_simulations. + eapply GiblePargenproof.transf_program_correct; eassumption. + apply forward_simulation_identity. } - split. auto. - apply forward_to_backward_simulation. - apply factor_forward_simulation. auto. eapply sd_traces. - eapply Verilog.semantics_determinate. - apply atomic_receptive. apply Cstrategy.semantics_strongly_receptive. - apply Verilog.semantics_determinate. -Qed. - -(*| -Backward Simulation -------------------- - -The following theorem is a *backward simulation* between the C and Verilog, -which proves the semantics preservation of the translation. We can assume that -the C and Verilog programs match, as we have proven previously in -``transf_hls_match`` that our translation implies that the ``match_prog`` -predicate will hold. -|*) - -Theorem c_semantic_preservation: - forall p tp, - match_prog p tp -> - backward_simulation (Csem.semantics p) (Verilog.semantics tp). -Proof. - intros. - apply compose_backward_simulation with (atomic (Cstrategy.semantics p)). - eapply sd_traces; eapply Verilog.semantics_determinate. - apply factor_backward_simulation. - apply Cstrategy.strategy_simulation. - apply Csem.semantics_single_events. - eapply ssr_well_behaved; eapply Cstrategy.semantics_strongly_receptive. - exact (proj2 (cstrategy_semantic_preservation _ _ H)). + auto. Qed. (*| @@ -507,35 +481,8 @@ function and that it succeeds. Theorem transf_c_program_correct: forall p tp, - transf_hls p = OK tp -> - backward_simulation (Csem.semantics p) (Verilog.semantics tp). -Proof. - intros. apply c_semantic_preservation. apply transf_hls_match; auto. -Qed. - -(*| -The final theorem of the semantic preservation of the translation of separate -translation units can also be proven correct, however, this is only because the -translation fails if more than one translation unit is passed to Vericert at the -moment. -|*) - -Theorem separate_transf_c_program_correct: - forall c_units verilog_units c_program, - nlist_forall2 (fun cu tcu => transf_hls cu = OK tcu) c_units verilog_units -> - link_list c_units = Some c_program -> - exists verilog_program, - link_list verilog_units = Some verilog_program - /\ backward_simulation (Csem.semantics c_program) - (Verilog.semantics verilog_program). + transf_hls_scheduled p = OK tp -> + forward_simulation (Cstrategy.semantics p) (GiblePar.GiblePar.semantics tp). Proof. - intros. - assert (nlist_forall2 match_prog c_units verilog_units). - { eapply nlist_forall2_imply. eauto. simpl; intros. - apply transf_hls_match; auto. } - assert (exists verilog_program, link_list verilog_units = Some verilog_program - /\ match_prog c_program verilog_program). - { eapply link_list_compose_passes; eauto. } - destruct H2 as (verilog_program & P & Q). - exists verilog_program; split; auto. apply c_semantic_preservation; auto. + intros. eapply cstrategy_semantic_preservation. apply transf_hls_match; auto. Qed. diff --git a/src/common/IntegerExtra.v b/src/common/IntegerExtra.v index 2b6cded..e56bda0 100644 --- a/src/common/IntegerExtra.v +++ b/src/common/IntegerExtra.v @@ -361,26 +361,6 @@ Module IntExtra. apply zlt_true. lia. lia. Qed. - Lemma bits_ofwords: - forall b4 b3 b2 b1 i, 0 <= i < zwordsize -> - testbit (ofbytes b4 b3 b2 b1) i = - if zlt i Byte.zwordsize - then Byte.testbit b1 i - else (if zlt i (2 * Byte.zwordsize) - then Byte.testbit b2 (i - Byte.zwordsize) - else (if zlt i (3 * Byte.zwordsize) - then Byte.testbit b2 (i - 2 * Byte.zwordsize) - else Byte.testbit b2 (i - 3 * Byte.zwordsize))). - Proof. - intros. unfold ofbytes. repeat (rewrite bits_or; auto). repeat (rewrite bits_shl; auto). - change (unsigned (repr Byte.zwordsize)) with Byte.zwordsize. - change (unsigned (repr (2 * Byte.zwordsize))) with (2 * Byte.zwordsize). - change (unsigned (repr (3 * Byte.zwordsize))) with (3 * Byte.zwordsize). - assert (zwordsize = 4 * Byte.zwordsize) by reflexivity. - destruct (zlt i Byte.zwordsize). - rewrite testbit_repr; auto. - Abort. - Lemma div_divs_equiv : forall x y, signed x >= 0 -> diff --git a/src/hls/CondElimproof.v b/src/hls/CondElimproof.v index 565adb1..7e6e035 100644 --- a/src/hls/CondElimproof.v +++ b/src/hls/CondElimproof.v @@ -310,6 +310,12 @@ Qed. Definition match_prog (p: program) (tp: program) := Linking.match_program (fun cu f tf => tf = transf_fundef f) eq p tp. +Lemma transf_program_match: + forall prog, match_prog prog (transf_program prog). +Proof. + intros. eapply Linking.match_transform_program_contextual. auto. +Qed. + Section CORRECTNESS. Context (prog tprog : program). diff --git a/src/hls/DMemorygen.v b/src/hls/DMemorygen.v index bef9f3d..d6c4fe8 100644 --- a/src/hls/DMemorygen.v +++ b/src/hls/DMemorygen.v @@ -2133,62 +2133,6 @@ Proof. apply Forall_forall. apply IHl. Qed. -(* Lemma transf_code_fold_correct: *) -(* forall l m state ram d' c' n, *) -(* fold_right (transf_maps state ram) (mod_datapath m, mod_controllogic m) l = (d', c') -> *) -(* Forall (fun x => x < n) (map fst l) -> *) -(* Forall (Pos.le n) (map snd l) -> *) -(* list_norepet (map fst l) -> *) -(* list_norepet (map snd l) -> *) -(* (forall p i c_s rs1 ar1 rs2 ar2 trs1 tar1 d_s, *) -(* i < n -> *) -(* all_match_empty_size m ar1 -> *) -(* all_match_empty_size m tar1 -> *) -(* match_module_to_ram m ram -> *) -(* (mod_datapath m)!i = Some d_s -> *) -(* (mod_controllogic m)!i = Some c_s -> *) -(* match_reg_assocs p rs1 trs1 -> *) -(* match_arr_assocs ar1 tar1 -> *) -(* max_reg_module m < p -> *) -(* exec_all d_s c_s rs1 ar1 rs2 ar2 -> *) -(* exists d_s' c_s' trs2 tar2, *) -(* d'!i = Some d_s' /\ c'!i = Some c_s' *) -(* /\ exec_all_ram ram d_s' c_s' trs1 tar1 trs2 tar2 *) -(* /\ match_reg_assocs p (merge_reg_assocs rs2) (merge_reg_assocs trs2) *) -(* /\ match_arr_assocs (merge_arr_assocs (ram_mem ram) (ram_size ram) ar2) *) -(* (merge_arr_assocs (ram_mem ram) (ram_size ram) tar2)). *) -(* Proof. *) -(* induction l as [| a l IHl]; simplify. *) -(* - match goal with *) -(* | H: (_, _) = (_, _) |- _ => inv H *) -(* end; *) -(* unfold exec_all in *; repeat inv_exists; simplify. *) -(* exploit match_states_same; *) -(* try match goal with *) -(* | H: stmnt_runp _ _ _ ?c _ _, H2: (mod_controllogic _) ! _ = Some ?c |- _ => apply H *) -(* end; eauto; mgen_crush; *) -(* try match goal with *) -(* | H: (mod_controllogic _) ! _ = Some ?c |- _ => *) -(* apply max_reg_stmnt_le_stmnt_tree in H; unfold max_reg_module in *; lia *) -(* end; intros; *) -(* exploit match_states_same; *) -(* try match goal with *) -(* | H: stmnt_runp _ _ _ ?c _ _, H2: (mod_datapath _) ! _ = Some ?c |- _ => apply H *) -(* end; eauto; mgen_crush; *) -(* try match goal with *) -(* | H: (mod_datapath _) ! _ = Some ?c |- _ => *) -(* apply max_reg_stmnt_le_stmnt_tree in H; unfold max_reg_module in *; lia *) -(* end; intros; *) -(* repeat match goal with *) -(* | |- exists _, _ => eexists *) -(* end; simplify; eauto; *) -(* unfold exec_all_ram; *) -(* repeat match goal with *) -(* | |- exists _, _ => eexists *) -(* end; simplify; eauto. *) -(* constructor. admit. *) -(* Abort. *) - Lemma empty_stack_transf : forall m, empty_stack (transf_module m) = empty_stack m. Proof. unfold empty_stack, transf_module; intros; repeat destruct_match; crush. Qed. diff --git a/src/hls/GiblePargenproof.v b/src/hls/GiblePargenproof.v index afa554c..95f0bb1 100644 --- a/src/hls/GiblePargenproof.v +++ b/src/hls/GiblePargenproof.v @@ -129,6 +129,12 @@ Definition state_lessdef := GiblePargenproofEquiv.match_states. Definition match_prog (prog : GibleSeq.program) (tprog : GiblePar.program) := match_program (fun cu f tf => transl_fundef f = OK tf) eq prog tprog. + Lemma transf_program_match: + forall p tp, transl_program p = OK tp -> match_prog p tp. + Proof. + intros. apply Linking.match_transform_partial_program; auto. + Qed. + (* TODO: Fix the `bb` and add matches for them. *) Inductive match_stackframes: GibleSeq.stackframe -> GiblePar.stackframe -> Prop := | match_stackframe: diff --git a/src/hls/GiblePargenproofCommon.v b/src/hls/GiblePargenproofCommon.v index 9254a3b..73bfa42 100644 --- a/src/hls/GiblePargenproofCommon.v +++ b/src/hls/GiblePargenproofCommon.v @@ -219,104 +219,6 @@ Qed. Definition all_mutexcl (f: forest) : Prop := forall x, predicated_mutexcl (f #r x). -(* Lemma map_mutexcl' : *) -(* forall A B (x: predicated A) (f: (pred_op * A) -> (pred_op * B)), *) -(* predicated_mutexcl x -> *) -(* (forall a pop, fst (f (pop, a)) ⇒ pop) -> *) -(* (forall a b, f a = f b -> a = b) -> *) -(* predicated_mutexcl (NE.map f x). *) -(* Proof. *) -(* induction x; cbn; intros. *) -(* - cbn. repeat constructor; intros. inv H2. inv H3. contradiction. *) -(* - constructor. intros. *) -(* assert (exists x0', x0 = f x0') by admit. *) -(* assert (exists y', y = f y') by admit. *) -(* simplify. unfold "⇒" in *; intros. *) -(* destruct x2, x1. *) -(* eapply H0 in H5. *) -(* inv H. *) -(* specialize (H6 (p, a0) (p0, a1)). *) -(* assert ((fst (p, a0)) ⇒ ¬ (fst (p0, a1))) by admit. *) -(* unfold "⇒" in H; simplify. *) -(* apply H in H5. *) - -Lemma map_deep_simplify : - forall A x (eqd: forall a b: pred_op * A, {a = b} + {a <> b}), - predicated_mutexcl x -> - predicated_mutexcl - (GiblePargenproofEquiv.NE.map (fun x : Predicate.pred_op * A => (deep_simplify peq (fst x), snd x)) x). -Proof. - intros. inv H. - assert (forall x0 y : pred_op * A, - GiblePargenproofEquiv.NE.In x0 - (GiblePargenproofEquiv.NE.map (fun x1 : Predicate.pred_op * A => (deep_simplify peq (fst x1), snd x1)) x) -> - GiblePargenproofEquiv.NE.In y - (GiblePargenproofEquiv.NE.map (fun x1 : Predicate.pred_op * A => (deep_simplify peq (fst x1), snd x1)) x) -> - x0 <> y -> fst x0 ⇒ ¬ fst y). - { intros. exploit NE.In_map2. eapply H. simplify. - exploit NE.In_map2. eapply H2. simplify. - specialize (H0 _ _ H4 H5). - destruct (eqd x1 x0); subst; try contradiction. - apply H0 in n. - unfold "⇒" in *; intros. rewrite negate_correct. - rewrite deep_simplify_correct. rewrite <- negate_correct. apply n. - erewrite <- deep_simplify_correct; eassumption. } - constructor; auto. - generalize dependent x. induction x; crush; [constructor|]. - inv H1. - assert (forall x0 y, GiblePargenproofEquiv.NE.In x0 x -> GiblePargenproofEquiv.NE.In y x -> x0 <> y -> fst x0 ⇒ ¬ fst y). - { intros. eapply H0; auto; constructor; tauto. } - assert (forall x0 y : pred_op * A, - GiblePargenproofEquiv.NE.In x0 - (GiblePargenproofEquiv.NE.map (fun x1 : Predicate.pred_op * A => (deep_simplify peq (fst x1), snd x1)) x) -> - GiblePargenproofEquiv.NE.In y - (GiblePargenproofEquiv.NE.map (fun x1 : Predicate.pred_op * A => (deep_simplify peq (fst x1), snd x1)) x) -> - x0 <> y -> fst x0 ⇒ ¬ fst y). - { intros. eapply H; auto; constructor; tauto. } - constructor; eauto. - unfold not; intros. - Abort. - -Lemma prune_predicated_mutexcl : - forall A (x: predicated A) y, - predicated_mutexcl x -> - prune_predicated x = Some y -> - predicated_mutexcl y. -Proof. - unfold prune_predicated; intros. - Abort. -(*#[local] Hint Resolve prune_predicated_mutexcl : mte.*) - -Lemma app_predicated_mutexcl : - forall A p (x: predicated A) y, - predicated_mutexcl x -> - predicated_mutexcl y -> - predicated_mutexcl (app_predicated p x y). -Proof. Abort. -(*#[local] Hint Resolve app_predicated_mutexcl : mte.*) - -Lemma seq_app_predicated_mutexcl : - forall A B (f: predicated (A -> B)) y, - predicated_mutexcl f -> - predicated_mutexcl y -> - predicated_mutexcl (seq_app f y). -Proof. Abort. -(* #[local] Hint Resolve seq_app_predicated_mutexcl : mte. *) - -Lemma all_predicated_mutexcl: - forall f l, - all_mutexcl f -> - predicated_mutexcl (merge (list_translation l f)). -Proof. Abort. -(* #[local] Hint Resolve all_predicated_mutexcl : mte. *) - -Lemma flap2_predicated_mutexcl : - forall A B C (f: predicated (A -> B -> C)) x y, - predicated_mutexcl f -> - predicated_mutexcl (flap2 f x y). -Proof. Abort. -(* #[local] Hint Resolve flap2_predicated_mutexcl : mte. *) - Lemma all_mutexcl_set : forall f n r, all_mutexcl f -> @@ -336,20 +238,3 @@ Proof. unfold pred_ret. repeat constructor; intros. inv H. inv H0. contradiction. Qed. #[local] Hint Resolve pred_ret_mutexcl : mte. - -Lemma all_mutexcl_maintained : - forall f p p' f' i, - all_mutexcl f -> - update (p, f) i = Some (p', f') -> - all_mutexcl f'. -Proof. - destruct i; cbn -[seq_app]; intros; unfold Option.bind in *; repeat destr; inv H0. - - trivial. - (* - apply prune_predicated_mutexcl in Heqo1; auto with mte. *) - (* - apply prune_predicated_mutexcl in Heqo0; auto with mte. *) - (* - apply prune_predicated_mutexcl in Heqo0; auto with mte. *) - (* apply app_predicated_mutexcl; auto with mte. *) - (* - unfold all_mutexcl; intros; rewrite forest_pred_reg; auto. *) - (* - unfold all_mutexcl; intros. unfold "<-e". unfold "#r". cbn. *) - (* fold (get_forest x f). auto. *) -Abort. diff --git a/src/hls/GiblePargenproofEquiv.v b/src/hls/GiblePargenproofEquiv.v index df14e31..6995bf5 100644 --- a/src/hls/GiblePargenproofEquiv.v +++ b/src/hls/GiblePargenproofEquiv.v @@ -422,59 +422,6 @@ Proof. constructor. Qed. -(* - -Lemma norm_expr_constant : - forall x m h t h' e p, - HN.norm_expression m x h = (t, h') -> - h ! e = Some p -> - h' ! e = Some p. -Proof. Abort. - -Definition sat_aequiv ap1 ap2 := - exists h p1 p2, - sat_equiv p1 p2 - /\ hash_predicate 1 ap1 h = (p1, h) - /\ hash_predicate 1 ap2 h = (p2, h). - -Lemma aequiv_symm : forall a b, sat_aequiv a b -> sat_aequiv b a. -Proof. - unfold sat_aequiv; simplify; do 3 eexists; simplify; [symmetry | |]; eauto. -Qed. - -Lemma existsh : - forall ap1, - exists h p1, - hash_predicate 1 ap1 h = (p1, h). -Proof. Abort. - -Lemma aequiv_refl : forall a, sat_aequiv a a. -Proof. - unfold sat_aequiv; intros. - pose proof (existsh a); simplify. - do 3 eexists; simplify; eauto. reflexivity. -Qed. - -Lemma aequiv_trans : - forall a b c, - sat_aequiv a b -> - sat_aequiv b c -> - sat_aequiv a c. -Proof. - unfold sat_aequiv; intros. - simplify. -Abort. - -Lemma norm_expr_mutexcl : - forall m pe h t h' e e' p p', - HN.norm_expression m pe h = (t, h') -> - predicated_mutexcl pe -> - t ! e = Some p -> - t ! e' = Some p' -> - e <> e' -> - p ⇒ ¬ p'. -Proof. Abort.*) - Definition pred_expr_eqb: forall pe1 pe2: pred_expr, {pe1 = pe2} + {pe1 <> pe2}. Proof. @@ -594,49 +541,6 @@ Definition ind_preds_l t := sat_predicate p1 c = true -> sat_predicate p2 c = false. -(*Lemma pred_equivalence_Some' : - forall ta tb e pe pe', - list_norepet (map fst ta) -> - list_norepet (map fst tb) -> - In (e, pe) ta -> - In (e, pe') tb -> - fold_right (fun p a => mk_pred_stmnt' (fst p) (snd p) ∧ a) T ta == - fold_right (fun p a => mk_pred_stmnt' (fst p) (snd p) ∧ a) T tb -> - pe == pe'. -Proof. - induction ta as [|hd tl Hta]; try solve [crush]. - - intros * NRP1 NRP2 IN1 IN2 FOLD. inv NRP1. inv IN1. - simpl in FOLD. - -Lemma pred_equivalence_Some : - forall (ta tb: PTree.t pred_op) e pe pe', - ta ! e = Some pe -> - tb ! e = Some pe' -> - mk_pred_stmnt ta == mk_pred_stmnt tb -> - pe == pe'. -Proof. - intros * SMEA SMEB EQ. unfold mk_pred_stmnt in *. - repeat rewrite PTree.fold_spec in EQ. - -Lemma pred_equivalence_None : - forall (ta tb: PTree.t pred_op) e pe, - (mk_pred_stmnt ta) == (mk_pred_stmnt tb) -> - ta ! e = Some pe -> - tb ! e = None -> - equiv pe ⟂. -Abort. - -Lemma pred_equivalence : - forall (ta tb: PTree.t pred_op) e pe, - equiv (mk_pred_stmnt ta) (mk_pred_stmnt tb) -> - ta ! e = Some pe -> - equiv pe ⟂ \/ (exists pe', tb ! e = Some pe' /\ equiv pe pe'). -Proof. - intros * EQ SME. destruct (tb ! e) eqn:HTB. - { right. econstructor. split; eauto. eapply pred_equivalence_Some; eauto. } - { left. eapply pred_equivalence_None; eauto. } -Qed.*) - Section CORRECT. Context {FUN TFUN: Type}. @@ -1163,20 +1067,6 @@ Module HashExprNorm(HS: Hashable). let (p2, h') := encode_expression 1%positive pe2 h in equiv_check p1 p2. - Lemma mk_pred_stmnt_equiv' : - forall l l' e p, - mk_pred_stmnt_l l == mk_pred_stmnt_l l' -> - In (e, p) l -> - list_norepet (map fst l) -> - (exists p', In (e, p') l' /\ p == p') - \/ ~ In e (map fst l') /\ p == ⟂. - Proof. Abort. - - Lemma mk_pred_stmnt_equiv : - forall tree tree', - mk_pred_stmnt tree == mk_pred_stmnt tree'. - Proof. Abort. - Definition tree_equiv_check_el eq_list (np2: PTree.t pred_op) (n: positive) (p: pred_op): bool := match np2 ! n with | Some p' => equiv_check_eq_list eq_list p p' @@ -1804,20 +1694,6 @@ Definition match_pred_states ht ! p = Some p_in -> sem_pred ctx p_in (pred_set !! p). -Lemma eval_hash_predicate : - forall max p_in ht p_out ht' br pred_set, - hash_predicate max p_in ht = (p_out, ht') -> - sem_pexpr ctx p_in br -> - match_pred_states ht' p_out pred_set -> - eval_predf pred_set p_out = br. -Proof. - induction p_in; simplify. - + repeat destruct_match. inv H. - unfold eval_predf. cbn. - inv H0. inv H4. unfold match_pred_states in H1. - specialize (H1 h br). -Abort. - Local Open Scope monad_scope. Fixpoint sem_valueF (e: expression) : option val := match e with diff --git a/src/hls/GiblePargenproofEvaluable.v b/src/hls/GiblePargenproofEvaluable.v index 3bb9c66..7e55909 100644 --- a/src/hls/GiblePargenproofEvaluable.v +++ b/src/hls/GiblePargenproofEvaluable.v @@ -274,40 +274,6 @@ Proof. eapply all_evaluable_cons in H. eauto. Qed. -Lemma all_evaluable_predicated_prod : - forall A B G (ctx: @ctx G) ps (a: predicated A) (b: predicated B), - all_evaluable ctx ps a -> - all_evaluable ctx ps b -> - all_evaluable ctx ps (predicated_prod a b). -Proof. - intros. unfold all_evaluable; intros. - unfold predicated_prod in *. exploit all_evaluable_map_and. - 2: { eauto. } - intros. 2: { eauto. } -Abort. (* Requires simple lemma to prove, but this lemma is not needed. *) - -Lemma cons_pred_expr_evaluable : - forall G (ctx: @ctx G) ps a b, - all_evaluable ctx ps a -> - all_evaluable ctx ps b -> - all_evaluable ctx ps (cons_pred_expr a b). -Proof. - unfold cons_pred_expr; intros. - apply all_evaluable_predicated_map. - (*now apply all_evaluable_predicated_prod.*) -Abort. - -Lemma fold_right_all_evaluable : - forall G (ctx: @ctx G) ps n, - Forall (all_evaluable ctx ps) (NE.to_list n) -> - all_evaluable ctx ps (NE.fold_right cons_pred_expr (NE.singleton (T, Enil)) n). -Proof. - induction n; cbn in *; intros. - - unfold cons_pred_expr. eapply all_evaluable_predicated_map. (*eapply all_evaluable_predicated_prod. - inv H. auto. unfold all_evaluable; intros. inv H0. exists true. constructor.*) admit. - - inv H. specialize (IHn H3). (*now eapply cons_pred_expr_evaluable.*) -Abort. - Lemma merge_all_evaluable : forall G (ctx: @ctx G) ps, pred_forest_evaluable ctx ps -> diff --git a/src/hls/GibleSeqgenproof.v b/src/hls/GibleSeqgenproof.v index 7451510..1d1d4b7 100644 --- a/src/hls/GibleSeqgenproof.v +++ b/src/hls/GibleSeqgenproof.v @@ -253,6 +253,12 @@ whole execution (in one big step) of the basic block. Definition match_prog (p: RTL.program) (tp: GibleSeq.program) := Linking.match_program (fun cu f tf => transl_fundef f = Errors.OK tf) eq p tp. + Lemma transf_program_match: + forall p tp, transl_program p = OK tp -> match_prog p tp. + Proof. + intros. apply Linking.match_transform_partial_program; auto. + Qed. + Context (TRANSL : match_prog prog tprog). Lemma symbols_preserved: diff --git a/src/hls/HTLgenproof.v b/src/hls/HTLgenproof.v deleted file mode 100644 index 1668580..0000000 --- a/src/hls/HTLgenproof.v +++ /dev/null @@ -1,2941 +0,0 @@ -(* - * Vericert: Verified high-level synthesis. - * Copyright (C) 2020 Yann Herklotz - * 2020 James Pollard - * - * 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 . - *) - -Require compcert.backend.RTL. -Require compcert.backend.Registers. -Require compcert.common.AST. -Require Import compcert.common.Globalenvs. -Require Import compcert.common.Linking. -Require Import compcert.common.Memory. -Require Import compcert.lib.Integers. - -Require Import vericert.common.IntegerExtra. -Require Import vericert.common.Vericertlib. -Require Import vericert.common.ZExtra. -Require Import vericert.hls.Array. -Require Import vericert.hls.AssocMap. -Require vericert.hls.HTL. -Require Import vericert.hls.HTLgen. -Require Import vericert.hls.HTLgenspec. -Require Import vericert.hls.ValueInt. -Require vericert.hls.Verilog. - -Require Import Lia. - -Local Open Scope assocmap. - -#[local] Hint Resolve Smallstep.forward_simulation_plus : htlproof. -#[local] Hint Resolve AssocMap.gss : htlproof. -#[local] Hint Resolve AssocMap.gso : htlproof. - -#[local] Hint Unfold find_assocmap AssocMapExt.get_default : htlproof. - -Inductive match_assocmaps : RTL.function -> RTL.regset -> assocmap -> Prop := - match_assocmap : forall f rs am, - (forall r, Ple r (RTL.max_reg_function f) -> - val_value_lessdef (Registers.Regmap.get r rs) am#r) -> - match_assocmaps f rs am. -#[local] Hint Constructors match_assocmaps : htlproof. - -Definition state_st_wf (m : HTL.module) (s : HTL.state) := - forall st asa asr res, - s = HTL.State res m st asa asr -> - asa!(m.(HTL.mod_st)) = Some (posToValue st). -#[local] Hint Unfold state_st_wf : htlproof. - -Inductive match_arrs (m : HTL.module) (f : RTL.function) (sp : Values.val) (mem : mem) : - Verilog.assocmap_arr -> Prop := -| match_arr : forall asa 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 (4 * ptr)))) - (Option.default (NToValue 0) - (Option.join (array_get_error (Z.to_nat ptr) stack)))) -> - match_arrs m f sp mem asa. - -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, 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) -> - stack_based (Option.default - Values.Vundef - (Mem.loadv AST.Mint32 m - (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_frames : list RTL.stackframe -> list HTL.stackframe -> Prop := -| match_frames_nil : - match_frames nil nil. - -Inductive match_constants : HTL.module -> assocmap -> Prop := - match_constant : - forall m asr, - asr!(HTL.mod_reset m) = Some (ZToValue 0) -> - asr!(HTL.mod_finish m) = Some (ZToValue 0) -> - match_constants m asr. - -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)) - (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) - (ASBP : arr_stack_based_pointers sp' mem (f.(RTL.fn_stacksize)) sp) - (BOUNDS : stack_bounds sp (f.(RTL.fn_stacksize)) mem) - (CONST : match_constants m asr), - match_states (RTL.State sf f sp st rs mem) - (HTL.State res m st asr asa) -| match_returnstate : - forall - v v' stack mem 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 : - forall f m m0 - (TF : tr_module f m), - match_states (RTL.Callstate nil (AST.Internal f) nil m0) (HTL.Callstate nil m nil). -#[local] 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 /\ - main_is_internal p = true. - -#[global] Instance TransfHTLLink (tr_fun: RTL.program -> Errors.res HTL.program): - TransfLink (fun (p1: RTL.program) (p2: HTL.program) => match_prog p1 p2). -Proof. - red. intros. exfalso. destruct (link_linkorder _ _ _ H) as [LO1 LO2]. - apply link_prog_inv in H. - - unfold match_prog in *. - unfold main_is_internal in *. simplify. repeat (unfold_match H4). - repeat (unfold_match H3). simplify. - subst. rewrite H0 in *. specialize (H (AST.prog_main p2)). - exploit H. - apply Genv.find_def_symbol. exists b. split. - assumption. apply Genv.find_funct_ptr_iff. eassumption. - apply Genv.find_def_symbol. exists b0. split. - assumption. apply Genv.find_funct_ptr_iff. eassumption. - intros. inv H3. inv H5. destruct H6. inv H5. -Qed. - -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. - -Lemma match_prog_matches : - forall p tp, match_prog p tp -> match_prog' p tp. -Proof. unfold match_prog. tauto. Qed. - -Lemma transf_program_match: - forall p tp, HTLgen.transl_program p = Errors.OK tp -> match_prog p tp. -Proof. - 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 : - forall f rs1 rs2 n v, - Plt (RTL.max_reg_function f) n -> - match_assocmaps f rs1 rs2 -> - match_assocmaps f rs1 (AssocMap.set n v rs2). -Proof. - inversion 2; subst. - intros. constructor. - intros. unfold find_assocmap. unfold AssocMapExt.get_default. - rewrite AssocMap.gso. eauto. - apply Pos.le_lt_trans with _ _ n in H2. - unfold not. intros. subst. eapply Pos.lt_irrefl. eassumption. assumption. -Qed. -#[local] Hint Resolve regs_lessdef_add_greater : htlproof. - -Lemma regs_lessdef_add_match : - forall f rs am r v v', - val_value_lessdef v v' -> - match_assocmaps f rs am -> - match_assocmaps f (Registers.Regmap.set r v rs) (AssocMap.set r v' am). -Proof. - inversion 2; subst. - constructor. intros. - destruct (peq r0 r); subst. - rewrite Registers.Regmap.gss. - unfold find_assocmap. unfold AssocMapExt.get_default. - rewrite AssocMap.gss. assumption. - - rewrite Registers.Regmap.gso; try assumption. - unfold find_assocmap. unfold AssocMapExt.get_default. - rewrite AssocMap.gso; eauto. -Qed. -#[local] 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; crush. - - symmetry. apply length_zero_iff_nil. auto. - - destruct l; crush. - rewrite list_repeat_cons. - crush. 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. - crush. - - rewrite <- (arr_wf a) in H. - apply list_combine_none. - 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; crush. - - rewrite nth_error_nil in H0. - discriminate. - - destruct l2 eqn:EQl2. crush. - simpl in H. inv H. - destruct n; simpl in *. - inv 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; crush; auto. - - destruct l2 eqn:EQl2. crush. - simpl in H. inv H. - destruct n; simpl in *. - inv 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. - -Ltac inv_state := - match goal with - MSTATE : match_states _ _ |- _ => - inversion MSTATE; - match goal with - TF : tr_module _ _ |- _ => - inversion TF; - match goal with - TC : forall _ _, - Maps.PTree.get _ _ = Some _ -> tr_code _ _ _ _ _ _ _ _ _, - H : Maps.PTree.get _ _ = Some _ |- _ => - apply TC in H; inversion H; - match goal with - TI : context[tr_instr] |- _ => - inversion TI - end - end - end -end; subst. - -Ltac unfold_func H := - match type of H with - | ?f = _ => unfold f in H; repeat (unfold_match H) - | ?f _ = _ => unfold f in H; repeat (unfold_match H) - | ?f _ _ = _ => unfold f in H; repeat (unfold_match H) - | ?f _ _ _ = _ => unfold f in H; repeat (unfold_match H) - | ?f _ _ _ _ = _ => unfold f in H; repeat (unfold_match H) - end. - -Lemma init_reg_assoc_empty : - forall f l, - match_assocmaps f (RTL.init_regs nil l) (HTL.init_regs nil l). -Proof. - induction l; simpl; constructor; intros. - - rewrite Registers.Regmap.gi. unfold find_assocmap. - unfold AssocMapExt.get_default. rewrite AssocMap.gempty. - constructor. - - - rewrite Registers.Regmap.gi. unfold find_assocmap. - unfold AssocMapExt.get_default. rewrite AssocMap.gempty. - 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. -#[local] Hint Resolve arr_lookup_some : htlproof. - -Section CORRECTNESS. - - Variable prog : RTL.program. - Variable tprog : HTL.program. - - Hypothesis TRANSL : match_prog prog tprog. - - 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. intros. eapply (Genv.find_symbol_match TRANSL'). Qed. - - Lemma function_ptr_translated: - forall (b: Values.block) (f: RTL.fundef), - Genv.find_funct_ptr ge b = Some f -> - 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 (cu & tf & P & Q & R); exists tf; auto. - Qed. - - Lemma functions_translated: - forall (v: Values.val) (f: RTL.fundef), - Genv.find_funct ge v = Some f -> - 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 (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'). - #[local] Hint Resolve senv_preserved : htlproof. - - Lemma ptrofs_inj : - forall a b, - Ptrofs.unsigned a = Ptrofs.unsigned b -> a = b. - Proof. - intros. rewrite <- Ptrofs.repr_unsigned. symmetry. rewrite <- Ptrofs.repr_unsigned. - rewrite H. auto. - Qed. - - Lemma op_stack_based : - forall F V sp v m args rs op ge pc' res0 pc f e fin rtrn st stk, - tr_instr fin rtrn st stk (RTL.Iop op args res0 pc') - (Verilog.Vnonblock (Verilog.Vvar res0) e) - (state_goto st pc') -> - reg_stack_based_pointers sp rs -> - (RTL.fn_code f) ! pc = Some (RTL.Iop op args res0 pc') -> - @Op.eval_operation F V ge (Values.Vptr sp Ptrofs.zero) op - (map (fun r : positive => Registers.Regmap.get r rs) args) m = Some v -> - stack_based v sp. - Proof. - Ltac solve_no_ptr := - match goal with - | H: reg_stack_based_pointers ?sp ?rs |- stack_based (Registers.Regmap.get ?r ?rs) _ => - solve [apply H] - | H1: reg_stack_based_pointers ?sp ?rs, H2: Registers.Regmap.get _ _ = Values.Vptr ?b ?i - |- context[Values.Vptr ?b _] => - let H := fresh "H" in - assert (H: stack_based (Values.Vptr b i) sp) by (rewrite <- H2; apply H1); simplify; solve [auto] - | |- context[Registers.Regmap.get ?lr ?lrs] => - destruct (Registers.Regmap.get lr lrs) eqn:?; simplify; auto - | |- stack_based (?f _) _ => unfold f - | |- stack_based (?f _ _) _ => unfold f - | |- stack_based (?f _ _ _) _ => unfold f - | |- stack_based (?f _ _ _ _) _ => unfold f - | H: ?f _ _ = Some _ |- _ => - unfold f in H; repeat (unfold_match H); inv H - | H: ?f _ _ _ _ _ _ = Some _ |- _ => - unfold f in H; repeat (unfold_match H); inv H - | H: map (fun r : positive => Registers.Regmap.get r _) ?args = _ |- _ => - destruct args; inv H - | |- context[if ?c then _ else _] => destruct c; try discriminate - | H: match _ with _ => _ end = Some _ |- _ => repeat (unfold_match H) - | H: match _ with _ => _ end = OK _ _ _ |- _ => repeat (unfold_match H) - | |- context[match ?g with _ => _ end] => destruct g; try discriminate - | |- _ => simplify; solve [auto] - end. - intros F V sp v m args rs op g pc' res0 pc f e fin rtrn st stk INSTR RSBP SEL EVAL. - inv INSTR. unfold translate_instr in H5. - unfold_match H5; repeat (unfold_match H5); repeat (simplify; solve_no_ptr). - Qed. - - Lemma int_inj : - forall x y, - Int.unsigned x = Int.unsigned y -> - x = y. - Proof. - intros. rewrite <- Int.repr_unsigned at 1. rewrite <- Int.repr_unsigned. - rewrite <- H. trivial. - Qed. - - Ltac eval_correct_tac := - match goal with - | |- context[valueToPtr] => unfold valueToPtr - | |- context[valueToInt] => unfold valueToInt - | |- context[bop] => unfold bop - | H : context[bop] |- _ => unfold bop in H - | |- context[boplit] => unfold boplit - | H : context[boplit] |- _ => unfold boplit in H - | |- context[boplitz] => unfold boplitz - | H : context[boplitz] |- _ => unfold boplitz in H - | |- val_value_lessdef Values.Vundef _ => solve [constructor] - | H : val_value_lessdef _ _ |- val_value_lessdef (Values.Vint _) _ => constructor; inv H - | |- val_value_lessdef (Values.Vint _) _ => constructor; auto - | H : ret _ _ = OK _ _ _ |- _ => inv H - | H : context[RTL.max_reg_function ?f] - |- context[_ (Registers.Regmap.get ?r ?rs) (Registers.Regmap.get ?r0 ?rs)] => - let HPle1 := fresh "HPle" in - let HPle2 := fresh "HPle" in - assert (HPle1 : Ple r (RTL.max_reg_function f)) by (eapply RTL.max_reg_function_use; eauto; simpl; auto); - assert (HPle2 : Ple r0 (RTL.max_reg_function f)) by (eapply RTL.max_reg_function_use; eauto; simpl; auto); - apply H in HPle1; apply H in HPle2; eexists; split; - [econstructor; eauto; constructor; trivial | inv HPle1; inv HPle2; try (constructor; auto)] - | H : context[RTL.max_reg_function ?f] - |- context[_ (Registers.Regmap.get ?r ?rs) _] => - let HPle1 := fresh "HPle" in - assert (HPle1 : Ple r (RTL.max_reg_function f)) by (eapply RTL.max_reg_function_use; eauto; simpl; auto); - apply H in HPle1; eexists; split; - [econstructor; eauto; constructor; trivial | inv HPle1; try (constructor; auto)] - | H : _ :: _ = _ :: _ |- _ => inv H - | |- context[match ?d with _ => _ end] => destruct d eqn:?; try discriminate - | H : match ?d with _ => _ end = _ |- _ => repeat unfold_match H - | H : match ?d with _ => _ end _ = _ |- _ => repeat unfold_match H - | |- Verilog.expr_runp _ _ _ ?f _ => - match f with - | Verilog.Vternary _ _ _ => idtac - | _ => econstructor - end - | |- val_value_lessdef (?f _ _) _ => unfold f - | |- val_value_lessdef (?f _) _ => unfold f - | H : ?f (Registers.Regmap.get _ _) _ = Some _ |- _ => - unfold f in H; repeat (unfold_match H) - | H1 : Registers.Regmap.get ?r ?rs = Values.Vint _, H2 : val_value_lessdef (Registers.Regmap.get ?r ?rs) _ - |- _ => rewrite H1 in H2; inv H2 - | |- _ => eexists; split; try constructor; solve [eauto] - | H : context[RTL.max_reg_function ?f] |- context[_ (Verilog.Vvar ?r) (Verilog.Vvar ?r0)] => - let HPle1 := fresh "H" in - let HPle2 := fresh "H" in - assert (HPle1 : Ple r (RTL.max_reg_function f)) by (eapply RTL.max_reg_function_use; eauto; simpl; auto); - assert (HPle2 : Ple r0 (RTL.max_reg_function f)) by (eapply RTL.max_reg_function_use; eauto; simpl; auto); - apply H in HPle1; apply H in HPle2; eexists; split; try constructor; eauto - | H : context[RTL.max_reg_function ?f] |- context[Verilog.Vvar ?r] => - let HPle := fresh "H" in - assert (HPle : Ple r (RTL.max_reg_function f)) by (eapply RTL.max_reg_function_use; eauto; simpl; auto); - apply H in HPle; eexists; split; try constructor; eauto - | |- context[if ?c then _ else _] => destruct c eqn:?; try discriminate - | H : ?b = _ |- _ = boolToValue ?b => rewrite H - end. - Ltac inv_lessdef := lazymatch goal with - | H2 : context[RTL.max_reg_function ?f], - H : Registers.Regmap.get ?r ?rs = _, - H1 : Registers.Regmap.get ?r0 ?rs = _ |- _ => - let HPle1 := fresh "HPle" in - assert (HPle1 : Ple r (RTL.max_reg_function f)) - by (eapply RTL.max_reg_function_use; eauto; simpl; auto); - apply H2 in HPle1; inv HPle1; - let HPle2 := fresh "HPle" in - assert (HPle2 : Ple r0 (RTL.max_reg_function f)) - by (eapply RTL.max_reg_function_use; eauto; simpl; auto); - apply H2 in HPle2; inv HPle2 - | H2 : context[RTL.max_reg_function ?f], H : Registers.Regmap.get ?r ?rs = _ |- _ => - let HPle1 := fresh "HPle" in - assert (HPle1 : Ple r (RTL.max_reg_function f)) - by (eapply RTL.max_reg_function_use; eauto; simpl; auto); - apply H2 in HPle1; inv HPle1 - end. - Ltac solve_cond := - match goal with - | H : context[match _ with _ => _ end] |- _ => repeat (unfold_merge H) - | H : ?f = _ |- context[boolToValue ?f] => rewrite H; solve [auto] - | H : Values.Vptr _ _ = Registers.Regmap.get ?r ?rs, - H2 : Registers.Regmap.get ?r ?rs = Values.Vint _ |- _ => - rewrite H2 in H; discriminate - | H : Values.Vundef = Registers.Regmap.get ?r ?rs, - H2 : Registers.Regmap.get ?r ?rs = Values.Vint _ |- _ => - rewrite H2 in H; discriminate - | H : Values.Vint _ = Registers.Regmap.get ?r ?rs, - H2 : Registers.Regmap.get ?r ?rs = Values.Vundef |- _ => - rewrite H2 in H; discriminate - | H : Values.Vint _ = Registers.Regmap.get ?r ?rs, - H2 : Registers.Regmap.get ?r ?rs = Values.Vtrue |- _ => - rewrite H2 in H; discriminate - | H : Values.Vint _ = Registers.Regmap.get ?r ?rs, - H2 : Registers.Regmap.get ?r ?rs = Values.Vfalse |- _ => - rewrite H2 in H; discriminate - | H : Values.Vint _ = Registers.Regmap.get ?r ?rs, - H2 : Registers.Regmap.get ?r ?rs = Values.Vptr _ _ |- _ => - rewrite H2 in H; discriminate - | H : Values.Vundef = Registers.Regmap.get ?r ?rs, - H2 : Registers.Regmap.get ?r ?rs = Values.Vptr _ _ |- _ => - rewrite H2 in H; discriminate - | H : Values.Vundef = Registers.Regmap.get ?r ?rs, - H2 : Registers.Regmap.get ?r ?rs = Values.Vtrue |- _ => - rewrite H2 in H; discriminate - | H : Values.Vundef = Registers.Regmap.get ?r ?rs, - H2 : Registers.Regmap.get ?r ?rs = Values.Vfalse |- _ => - rewrite H2 in H; discriminate - | H : Values.Vptr _ _ = Registers.Regmap.get ?r ?rs, - H2 : Registers.Regmap.get ?r ?rs = Values.Vundef |- _ => - rewrite H2 in H; discriminate - | H : Values.Vptr _ _ = Registers.Regmap.get ?r ?rs, - H2 : Registers.Regmap.get ?r ?rs = Values.Vtrue |- _ => - rewrite H2 in H; discriminate - | H : Values.Vptr _ _ = Registers.Regmap.get ?r ?rs, - H2 : Registers.Regmap.get ?r ?rs = Values.Vfalse |- _ => - rewrite H2 in H; discriminate - | |- context[val_value_lessdef Values.Vundef _] => - econstructor; split; econstructor; econstructor; auto; solve [constructor] - | H1 : Registers.Regmap.get ?r ?rs = Values.Vint _, - H2 : Values.Vint _ = Registers.Regmap.get ?r ?rs, - H3 : Registers.Regmap.get ?r0 ?rs = Values.Vint _, - H4 : Values.Vint _ = Registers.Regmap.get ?r0 ?rs|- _ => - rewrite H1 in H2; rewrite H3 in H4; inv H2; inv H4; unfold valueToInt in *; constructor - | H1 : Registers.Regmap.get ?r ?rs = Values.Vptr _ _, - H2 : Values.Vptr _ _ = Registers.Regmap.get ?r ?rs, - H3 : Registers.Regmap.get ?r0 ?rs = Values.Vptr _ _, - H4 : Values.Vptr _ _ = Registers.Regmap.get ?r0 ?rs|- _ => - rewrite H1 in H2; rewrite H3 in H4; inv H2; inv H4; unfold valueToInt in *; constructor; - unfold Ptrofs.ltu, Int.ltu in *; unfold Ptrofs.of_int in *; - repeat (rewrite Ptrofs.unsigned_repr in *; auto using Int.unsigned_range_2) - | H : _ :: _ = _ :: _ |- _ => inv H - | H : ret _ _ = OK _ _ _ |- _ => inv H - | |- _ => - eexists; split; [ econstructor; econstructor; auto - | simplify; inv_lessdef; repeat (unfold valueToPtr, valueToInt in *; solve_cond); - unfold valueToPtr in * - ] - end. - - Lemma eval_cond_correct : - forall stk f sp pc rs m res ml st asr asa e b f' s s' args i cond, - match_states (RTL.State stk f sp pc rs m) (HTL.State res ml st asr asa) -> - (forall v, In v args -> Ple v (RTL.max_reg_function f)) -> - Op.eval_condition cond (map (fun r : positive => Registers.Regmap.get r rs) args) m = Some b -> - translate_condition cond args s = OK e s' i -> - Verilog.expr_runp f' asr asa e (boolToValue b). - Proof. - intros stk f sp pc rs m res ml st asr asa e b f' s s' args i cond MSTATE MAX_FUN EVAL TR_INSTR. - pose proof MSTATE as MSTATE_2. inv MSTATE. - inv MASSOC. unfold translate_condition, translate_comparison, - translate_comparisonu, translate_comparison_imm, - translate_comparison_immu in TR_INSTR; - repeat unfold_match TR_INSTR; try inv TR_INSTR; simplify_val; - unfold Values.Val.cmp_bool, Values.Val.of_optbool, bop, Values.Val.cmpu_bool, - Int.cmpu in *; - repeat unfold_match EVAL. - - repeat econstructor. repeat unfold_match Heqo. simplify_val. - pose proof (MAX_FUN p) as MAX_FUN_P. apply H in MAX_FUN_P; auto. - pose proof (MAX_FUN p0) as MAX_FUN_P0. apply H in MAX_FUN_P0; auto. - inv MAX_FUN_P; inv MAX_FUN_P0; solve_cond. - - repeat econstructor. repeat unfold_match Heqo. simplify_val. - pose proof (MAX_FUN p) as MAX_FUN_P. apply H in MAX_FUN_P; auto. - pose proof (MAX_FUN p0) as MAX_FUN_P0. apply H in MAX_FUN_P0; auto. - inv MAX_FUN_P; inv MAX_FUN_P0; solve_cond. - - repeat econstructor. repeat unfold_match Heqo. simplify_val. - pose proof (MAX_FUN p) as MAX_FUN_P. apply H in MAX_FUN_P; auto. - pose proof (MAX_FUN p0) as MAX_FUN_P0. apply H in MAX_FUN_P0; auto. - inv MAX_FUN_P; inv MAX_FUN_P0; solve_cond. - - repeat econstructor. repeat unfold_match Heqo. simplify_val. - pose proof (MAX_FUN p) as MAX_FUN_P. apply H in MAX_FUN_P; auto. - pose proof (MAX_FUN p0) as MAX_FUN_P0. apply H in MAX_FUN_P0; auto. - inv MAX_FUN_P; inv MAX_FUN_P0; solve_cond. - - repeat econstructor. repeat unfold_match Heqo. simplify_val. - pose proof (MAX_FUN p) as MAX_FUN_P. apply H in MAX_FUN_P; auto. - pose proof (MAX_FUN p0) as MAX_FUN_P0. apply H in MAX_FUN_P0; auto. - inv MAX_FUN_P; inv MAX_FUN_P0; solve_cond. - - repeat econstructor. repeat unfold_match Heqo. simplify_val. - pose proof (MAX_FUN p) as MAX_FUN_P. apply H in MAX_FUN_P; auto. - pose proof (MAX_FUN p0) as MAX_FUN_P0. apply H in MAX_FUN_P0; auto. - inv MAX_FUN_P; inv MAX_FUN_P0; solve_cond. - - repeat econstructor. repeat unfold_match Heqo; simplify_val. - pose proof (MAX_FUN p) as MAX_FUN_P. apply H in MAX_FUN_P; auto. - pose proof (MAX_FUN p0) as MAX_FUN_P0. apply H in MAX_FUN_P0; auto. - inv MAX_FUN_P; inv MAX_FUN_P0; solve_cond. - - repeat econstructor. repeat unfold_match Heqo; simplify_val. - pose proof (MAX_FUN p) as MAX_FUN_P. apply H in MAX_FUN_P; auto. - pose proof (MAX_FUN p0) as MAX_FUN_P0. apply H in MAX_FUN_P0; auto. - inv MAX_FUN_P; inv MAX_FUN_P0; try solve_cond. simplify_val. - rewrite Heqv0 in H3. rewrite Heqv in H2. inv H2. inv H3. - unfold Ptrofs.ltu. unfold Int.ltu. - rewrite Ptrofs.unsigned_repr by apply Int.unsigned_range_2. - rewrite Ptrofs.unsigned_repr by apply Int.unsigned_range_2. auto. - - repeat econstructor. unfold Verilog.binop_run. - pose proof (MAX_FUN p) as MAX_FUN_P. apply H in MAX_FUN_P; auto. - pose proof (MAX_FUN p0) as MAX_FUN_P0. apply H in MAX_FUN_P0; auto. - inv MAX_FUN_P; inv MAX_FUN_P0; simplify_val; solve_cond. - - repeat econstructor. simplify_val. - pose proof (MAX_FUN p) as MAX_FUN_P. apply H in MAX_FUN_P; auto. - pose proof (MAX_FUN p0) as MAX_FUN_P0. apply H in MAX_FUN_P0; auto. - inv MAX_FUN_P; inv MAX_FUN_P0; try solve_cond. simplify_val. - rewrite Heqv0 in H3. rewrite Heqv in H2. inv H2. inv H3. - unfold Ptrofs.ltu. unfold Int.ltu. - rewrite Ptrofs.unsigned_repr by apply Int.unsigned_range_2. - rewrite Ptrofs.unsigned_repr by apply Int.unsigned_range_2. auto. - - repeat econstructor. unfold Verilog.binop_run. - pose proof (MAX_FUN p) as MAX_FUN_P. apply H in MAX_FUN_P; auto. - pose proof (MAX_FUN p0) as MAX_FUN_P0. apply H in MAX_FUN_P0; auto. - inv MAX_FUN_P; inv MAX_FUN_P0; simplify_val; solve_cond. - - repeat econstructor. simplify_val. - pose proof (MAX_FUN p) as MAX_FUN_P. apply H in MAX_FUN_P; auto. - pose proof (MAX_FUN p0) as MAX_FUN_P0. apply H in MAX_FUN_P0; auto. - inv MAX_FUN_P; inv MAX_FUN_P0; try solve_cond. simplify_val. - rewrite Heqv0 in H3. rewrite Heqv in H2. inv H2. inv H3. - unfold Ptrofs.ltu. unfold Int.ltu. - rewrite Ptrofs.unsigned_repr by apply Int.unsigned_range_2. - rewrite Ptrofs.unsigned_repr by apply Int.unsigned_range_2. auto. - - repeat econstructor. unfold Verilog.binop_run. - pose proof (MAX_FUN p) as MAX_FUN_P. apply H in MAX_FUN_P; auto. - pose proof (MAX_FUN p0) as MAX_FUN_P0. apply H in MAX_FUN_P0; auto. - inv MAX_FUN_P; inv MAX_FUN_P0; simplify_val; solve_cond. - - repeat econstructor. simplify_val. - pose proof (MAX_FUN p) as MAX_FUN_P. apply H in MAX_FUN_P; auto. - pose proof (MAX_FUN p0) as MAX_FUN_P0. apply H in MAX_FUN_P0; auto. - inv MAX_FUN_P; inv MAX_FUN_P0; try solve_cond. simplify_val. - rewrite Heqv0 in H3. rewrite Heqv in H2. inv H2. inv H3. - unfold Ptrofs.ltu. unfold Int.ltu. - rewrite Ptrofs.unsigned_repr by apply Int.unsigned_range_2. - rewrite Ptrofs.unsigned_repr by apply Int.unsigned_range_2. auto. - - repeat econstructor. simplify_val. - pose proof (MAX_FUN p) as MAX_FUN_P. apply H in MAX_FUN_P; auto. - inv MAX_FUN_P; simplify_val; try solve_cond. - rewrite Heqv in H0. inv H0. auto. - - repeat econstructor. simplify_val. - pose proof (MAX_FUN p) as MAX_FUN_P. apply H in MAX_FUN_P; auto. - inv MAX_FUN_P; simplify_val; try solve_cond. - rewrite Heqv in H0. inv H0. auto. - - repeat econstructor. simplify_val. - pose proof (MAX_FUN p) as MAX_FUN_P. apply H in MAX_FUN_P; auto. - inv MAX_FUN_P; simplify_val; try solve_cond. - rewrite Heqv in H0. inv H0. auto. - - repeat econstructor. simplify_val. - pose proof (MAX_FUN p) as MAX_FUN_P. apply H in MAX_FUN_P; auto. - inv MAX_FUN_P; simplify_val; try solve_cond. - rewrite Heqv in H0. inv H0. auto. - - repeat econstructor. simplify_val. - pose proof (MAX_FUN p) as MAX_FUN_P. apply H in MAX_FUN_P; auto. - inv MAX_FUN_P; simplify_val; try solve_cond. - rewrite Heqv in H0. inv H0. auto. - - repeat econstructor. simplify_val. - pose proof (MAX_FUN p) as MAX_FUN_P. apply H in MAX_FUN_P; auto. - inv MAX_FUN_P; simplify_val; try solve_cond. - rewrite Heqv in H0. inv H0. auto. - - repeat econstructor. simplify_val. - pose proof (MAX_FUN p) as MAX_FUN_P. apply H in MAX_FUN_P; auto. - inv MAX_FUN_P; simplify_val; try solve_cond. - rewrite Heqv in H0. inv H0. auto. - - repeat econstructor. simplify_val. - pose proof (MAX_FUN p) as MAX_FUN_P. apply H in MAX_FUN_P; auto. - inv MAX_FUN_P; simplify_val; try solve_cond. - rewrite Heqv in H0. inv H0. auto. - - repeat econstructor. simplify_val. - pose proof (MAX_FUN p) as MAX_FUN_P. apply H in MAX_FUN_P; auto. - inv MAX_FUN_P; simplify_val; try solve_cond. - rewrite Heqv in H0. inv H0. auto. - - repeat econstructor. simplify_val. - pose proof (MAX_FUN p) as MAX_FUN_P. apply H in MAX_FUN_P; auto. - inv MAX_FUN_P; simplify_val; try solve_cond. - rewrite Heqv in H0. inv H0. auto. - Qed. - - Lemma eval_cond_correct' : - forall e stk f sp pc rs m res ml st asr asa v f' s s' args i cond, - match_states (RTL.State stk f sp pc rs m) (HTL.State res ml st asr asa) -> - (forall v, In v args -> Ple v (RTL.max_reg_function f)) -> - Values.Val.of_optbool None = v -> - translate_condition cond args s = OK e s' i -> - exists v', Verilog.expr_runp f' asr asa e v' /\ val_value_lessdef v v'. - intros e stk f sp pc rs m res ml st asr asa v f' s s' args i cond MSTATE MAX_FUN EVAL TR_INSTR. - unfold translate_condition, translate_comparison, translate_comparisonu, - translate_comparison_imm, translate_comparison_immu, bop, boplit in *. - repeat unfold_match TR_INSTR; inv TR_INSTR; repeat econstructor. - Qed. - - Lemma eval_correct_Oshrximm : - forall s sp rs m v e asr asa f f' stk s' i pc res0 pc' args res ml st n, - match_states (RTL.State stk f sp pc rs m) (HTL.State res ml st asr asa) -> - (RTL.fn_code f) ! pc = Some (RTL.Iop (Op.Oshrximm n) args res0 pc') -> - Op.eval_operation ge sp (Op.Oshrximm n) - (List.map (fun r : BinNums.positive => - Registers.Regmap.get r rs) args) m = Some v -> - translate_instr (Op.Oshrximm n) args s = OK e s' i -> - exists v', Verilog.expr_runp f' asr asa e v' /\ val_value_lessdef v v'. - Proof. - intros s sp rs m v e asr asa f f' stk s' i pc pc' res0 args res ml st n MSTATE INSTR EVAL TR_INSTR. - pose proof MSTATE as MSTATE_2. inv MSTATE. - inv MASSOC. unfold translate_instr in TR_INSTR; repeat (unfold_match TR_INSTR); inv TR_INSTR; - unfold Op.eval_operation in EVAL; repeat (unfold_match EVAL); inv EVAL. - (*repeat (simplify; eval_correct_tac; unfold valueToInt in * ). - destruct (Z_lt_ge_dec (Int.signed i0) 0). - econstructor.*) - unfold Values.Val.shrx in *. - destruct v0; try discriminate. - destruct (Int.ltu n (Int.repr 31)) eqn:?; try discriminate. - inversion H1. clear H1. - assert (Int.unsigned n <= 30). - { unfold Int.ltu in *. destruct (zlt (Int.unsigned n) (Int.unsigned (Int.repr 31))); try discriminate. - rewrite Int.unsigned_repr in l by (simplify; lia). - replace 31 with (Z.succ 30) in l by auto. - apply Zlt_succ_le in l. - auto. - } - rewrite IntExtra.shrx_shrx_alt_equiv in H2 by auto. - unfold IntExtra.shrx_alt in *. - destruct (zlt (Int.signed i0) 0). - - repeat econstructor; unfold valueToBool, boolToValue, uvalueToZ, natToValue; - repeat (simplify; eval_correct_tac). - inv_lessdef. unfold valueToInt in *. rewrite H3 in H1. - inv H1. - unfold Int.lt in *. rewrite zlt_true in Heqb0. simplify. - rewrite Int.unsigned_repr in Heqb0. discriminate. - simplify; lia. - unfold ZToValue. rewrite Int.signed_repr by (simplify; lia). - auto. - rewrite H3 in H1; discriminate. - rewrite H3 in H2; discriminate. - rewrite IntExtra.shrx_shrx_alt_equiv; auto. unfold IntExtra.shrx_alt. rewrite zlt_true; try lia. - simplify. inv_lessdef. unfold valueToInt in *. - rewrite H3 in H1. auto. inv H1. auto. - rewrite H3 in H1. discriminate. - rewrite H3 in H2. discriminate. - - econstructor; econstructor; [eapply Verilog.erun_Vternary_false|idtac]; repeat econstructor; unfold valueToBool, boolToValue, uvalueToZ, natToValue; - repeat (simplify; eval_correct_tac). - inv_lessdef. unfold valueToInt in *. rewrite H3 in H1. - inv H1. - unfold Int.lt in *. rewrite zlt_false in Heqb0. simplify. - rewrite Int.unsigned_repr in Heqb0. lia. - simplify; lia. - unfold ZToValue. rewrite Int.signed_repr by (simplify; lia). - auto. - rewrite H3 in H1; discriminate. - rewrite H3 in H2; discriminate. - rewrite IntExtra.shrx_shrx_alt_equiv; auto. unfold IntExtra.shrx_alt. rewrite zlt_false; try lia. - simplify. inv_lessdef. unfold valueToInt in *. - rewrite H3 in H1. auto. inv H1. auto. - rewrite H3 in H1. discriminate. - rewrite H3 in H2. discriminate. - Qed. - - Lemma eval_correct : - forall s sp op rs m v e asr asa f f' stk s' i pc res0 pc' args res ml st, - match_states (RTL.State stk f sp pc rs m) (HTL.State res ml st asr asa) -> - (RTL.fn_code f) ! pc = Some (RTL.Iop op args res0 pc') -> - Op.eval_operation ge sp op - (List.map (fun r : BinNums.positive => Registers.Regmap.get r rs) args) m = Some v -> - translate_instr op args s = OK e s' i -> - exists v', Verilog.expr_runp f' asr asa e v' /\ val_value_lessdef v v'. - Proof. - intros s sp op rs m v e asr asa f f' stk s' i pc pc' res0 args res ml st MSTATE INSTR EVAL TR_INSTR. - pose proof MSTATE as MSTATE_2. inv MSTATE. - inv MASSOC. unfold translate_instr in TR_INSTR; repeat (unfold_match TR_INSTR); inv TR_INSTR; - unfold Op.eval_operation in EVAL; repeat (unfold_match EVAL); inv EVAL; - repeat (simplify; eval_correct_tac; unfold valueToInt in *). - - pose proof Integers.Ptrofs.agree32_sub as H2; unfold Integers.Ptrofs.agree32 in H2. - unfold Ptrofs.of_int. simpl. - apply ptrofs_inj. assert (Archi.ptr64 = false) by auto. eapply H2 in H3. - rewrite Ptrofs.unsigned_repr. apply H3. replace Ptrofs.max_unsigned with Int.max_unsigned; auto. - apply Int.unsigned_range_2. - auto. rewrite Ptrofs.unsigned_repr. replace Ptrofs.max_unsigned with Int.max_unsigned; auto. - apply Int.unsigned_range_2. rewrite Ptrofs.unsigned_repr. auto. - replace Ptrofs.max_unsigned with Int.max_unsigned; auto. - apply Int.unsigned_range_2. - - pose proof Integers.Ptrofs.agree32_sub as AGR; unfold Integers.Ptrofs.agree32 in AGR. - assert (ARCH: Archi.ptr64 = false) by auto. eapply AGR in ARCH. - apply int_inj. unfold Ptrofs.to_int. rewrite Int.unsigned_repr. - apply ARCH. pose proof Ptrofs.unsigned_range_2. - replace Ptrofs.max_unsigned with Int.max_unsigned; auto. - pose proof Ptrofs.agree32_of_int. unfold Ptrofs.agree32 in H2. - eapply H2 in ARCH. apply ARCH. - pose proof Ptrofs.agree32_of_int. unfold Ptrofs.agree32 in H2. - eapply H2 in ARCH. apply ARCH. - - rewrite H0 in Heqb. rewrite H1 in Heqb. discriminate. - - rewrite Heqb in Heqb0. discriminate. - - rewrite H0 in Heqb. rewrite H1 in Heqb. discriminate. - - rewrite Heqb in Heqb0. discriminate. - (*- unfold Int.ror. unfold Int.or. unfold Int.shru, Int.shl, Int.sub. unfold intToValue. unfold Int.modu, - repeat (rewrite Int.unsigned_repr). auto.*) - - assert (Int.unsigned n <= 30). - { unfold Int.ltu in *. destruct (zlt (Int.unsigned n) (Int.unsigned (Int.repr 31))); try discriminate. - rewrite Int.unsigned_repr in l by (simplify; lia). - replace 31 with (Z.succ 30) in l by auto. - apply Zlt_succ_le in l. - auto. - } - destruct (zlt (Int.signed i0) 0). - + repeat econstructor; unfold valueToBool, boolToValue, uvalueToZ, natToValue; - repeat (simplify; eval_correct_tac). - rewrite IntExtra.shrx_shrx_alt_equiv; auto. unfold IntExtra.shrx_alt. rewrite zlt_true; try lia. - simplify. inv_lessdef. unfold valueToInt in *. - rewrite Heqv0 in H0. auto. inv H0. auto. - rewrite Heqv0 in H2. discriminate. - unfold valueToInt in l. auto. - inv_lessdef. unfold valueToInt in *. rewrite Heqv0 in H0. - inv H0. - unfold Int.lt in *. rewrite zlt_true in Heqb0. simplify. - rewrite Int.unsigned_repr in Heqb0. discriminate. - simplify; lia. - unfold ZToValue. rewrite Int.signed_repr by (simplify; lia). - auto. - rewrite Heqv0 in H0; discriminate. - rewrite Heqv0 in H2; discriminate. - + eapply Verilog.erun_Vternary_false; repeat econstructor; unfold valueToBool, boolToValue, uvalueToZ, natToValue; - repeat (simplify; eval_correct_tac). - rewrite IntExtra.shrx_shrx_alt_equiv; auto. unfold IntExtra.shrx_alt. rewrite zlt_false; try lia. - simplify. inv_lessdef. unfold valueToInt in *. - rewrite Heqv0 in H0. auto. inv H0. auto. - rewrite Heqv0 in H2. discriminate. - unfold valueToInt in *. auto. - inv_lessdef. unfold valueToInt in *. - rewrite Heqv0 in H0. - inv H0. - unfold Int.lt in *. rewrite zlt_false in Heqb0. simplify. - rewrite Int.unsigned_repr in Heqb0. lia. - simplify; lia. - unfold ZToValue. rewrite Int.signed_repr by (simplify; lia). - auto. - rewrite Heqv0 in H0; discriminate. - rewrite Heqv0 in H2; discriminate. - - unfold Op.eval_addressing32 in *. repeat (unfold_match H2); inv H2. - + unfold translate_eff_addressing in *. repeat (unfold_match H1). - destruct v0; inv Heql; rewrite H2; inv H1; repeat eval_correct_tac. - pose proof Integers.Ptrofs.agree32_add as AGR; unfold Integers.Ptrofs.agree32 in AGR. unfold ZToValue. - apply ptrofs_inj. unfold Ptrofs.of_int. rewrite Ptrofs.unsigned_repr. - apply AGR. auto. rewrite H2 in H0. inv H0. unfold valueToPtr. unfold Ptrofs.of_int. - rewrite Ptrofs.unsigned_repr. auto. replace Ptrofs.max_unsigned with Int.max_unsigned by auto. - apply Int.unsigned_range_2. - rewrite Ptrofs.unsigned_repr. auto. replace Ptrofs.max_unsigned with Int.max_unsigned by auto. - apply Int.unsigned_range_2. - replace Ptrofs.max_unsigned with Int.max_unsigned by auto. - apply Int.unsigned_range_2. - + unfold translate_eff_addressing in *. repeat (unfold_match H1). inv H1. - inv Heql. unfold boplitz. repeat (simplify; eval_correct_tac). - all: repeat (unfold_match Heqv). - * inv Heqv. unfold valueToInt in *. inv H2. inv H0. unfold valueToInt in *. trivial. - * constructor. unfold valueToPtr, ZToValue in *. - pose proof Integers.Ptrofs.agree32_add as AGR; unfold Integers.Ptrofs.agree32 in AGR. unfold ZToValue. - apply ptrofs_inj. unfold Ptrofs.of_int. rewrite Ptrofs.unsigned_repr. - apply AGR. auto. inv Heqv. rewrite Int.add_commut. - apply AGR. auto. inv H1. inv H0. unfold valueToPtr. unfold Ptrofs.of_int. - rewrite Ptrofs.unsigned_repr. auto. replace Ptrofs.max_unsigned with Int.max_unsigned by auto. - apply Int.unsigned_range_2. - unfold Ptrofs.of_int. - rewrite Ptrofs.unsigned_repr. inv H0. auto. replace Ptrofs.max_unsigned with Int.max_unsigned by auto. - apply Int.unsigned_range_2. - rewrite Ptrofs.unsigned_repr. auto. replace Ptrofs.max_unsigned with Int.max_unsigned by auto. - apply Int.unsigned_range_2. - apply Int.unsigned_range_2. - * constructor. unfold valueToPtr, ZToValue in *. - pose proof Integers.Ptrofs.agree32_add as AGR; unfold Integers.Ptrofs.agree32 in AGR. unfold ZToValue. - apply ptrofs_inj. unfold Ptrofs.of_int. rewrite Ptrofs.unsigned_repr. - apply AGR. auto. inv Heqv. - apply AGR. auto. inv H0. unfold valueToPtr, Ptrofs.of_int. rewrite Ptrofs.unsigned_repr. auto. - replace Ptrofs.max_unsigned with Int.max_unsigned by auto. - apply Int.unsigned_range_2. - inv H1. unfold valueToPtr, Ptrofs.of_int. rewrite Ptrofs.unsigned_repr. auto. - replace Ptrofs.max_unsigned with Int.max_unsigned by auto. - apply Int.unsigned_range_2. - rewrite Ptrofs.unsigned_repr. auto. - replace Ptrofs.max_unsigned with Int.max_unsigned by auto. - apply Int.unsigned_range_2. apply Int.unsigned_range_2. - + unfold translate_eff_addressing in *. repeat (unfold_match H1). inv H1. - inv Heql. unfold boplitz. repeat (simplify; eval_correct_tac). - all: repeat (unfold_match Heqv). - * unfold Values.Val.mul in Heqv. repeat (unfold_match Heqv). inv Heqv. inv H3. - unfold valueToInt, ZToValue. auto. - * unfold Values.Val.mul in Heqv. repeat (unfold_match Heqv). - * unfold Values.Val.mul in Heqv. repeat (unfold_match Heqv). - * constructor. unfold valueToPtr, ZToValue. unfold Values.Val.mul in Heqv. repeat (unfold_match Heqv). - + unfold translate_eff_addressing in *. repeat (unfold_match H1). inv H1. - inv Heql. unfold boplitz. repeat (simplify; eval_correct_tac). - all: repeat (unfold_match Heqv). - unfold valueToPtr, ZToValue. - repeat unfold_match Heqv0. unfold Values.Val.mul in Heqv1. repeat unfold_match Heqv1. - inv Heqv1. inv Heqv0. unfold valueToInt in *. - assert (HPle1 : Ple r0 (RTL.max_reg_function f)) by (eapply RTL.max_reg_function_use; eauto; simpl; auto). - apply H in HPle1. inv HPle1. unfold valueToInt in *. rewrite Heqv2 in H2. inv H2. auto. - rewrite Heqv2 in H2. inv H2. - rewrite Heqv2 in H3. discriminate. - repeat unfold_match Heqv0. unfold Values.Val.mul in Heqv1. repeat unfold_match Heqv1. - repeat unfold_match Heqv0. unfold Values.Val.mul in Heqv1. repeat unfold_match Heqv1. - constructor. unfold valueToPtr, ZToValue. inv Heqv0. inv Heqv1. - assert (HPle1 : Ple r0 (RTL.max_reg_function f)) by (eapply RTL.max_reg_function_use; eauto; simpl; auto). - apply H in HPle1. inv HPle1. unfold valueToInt in *. rewrite Heqv2 in H3. inv H3. - - pose proof Integers.Ptrofs.agree32_add as AGR; unfold Integers.Ptrofs.agree32 in AGR. unfold ZToValue. - apply ptrofs_inj. unfold Ptrofs.of_int. rewrite Ptrofs.unsigned_repr. - apply AGR. auto. inv H2. unfold valueToPtr, Ptrofs.of_int. rewrite Ptrofs.unsigned_repr. auto. - replace Ptrofs.max_unsigned with Int.max_unsigned by auto. apply Int.unsigned_range_2. - apply Ptrofs.unsigned_repr. apply Int.unsigned_range_2. apply Int.unsigned_range_2. - - rewrite Heqv2 in H3. inv H3. - - rewrite Heqv2 in H4. inv H4. - + unfold translate_eff_addressing in *. repeat (unfold_match H1). inv H1. - inv Heql. unfold boplitz. repeat (simplify; eval_correct_tac). - all: repeat (unfold_match Heqv). - eexists. split. constructor. - constructor. unfold valueToPtr, ZToValue. rewrite Ptrofs.add_zero_l. unfold Ptrofs.of_int. - rewrite Int.unsigned_repr. symmetry. apply Ptrofs.repr_unsigned. - unfold check_address_parameter_unsigned in *. apply Ptrofs.unsigned_range_2. - - destruct (Op.eval_condition cond (map (fun r : positive => Registers.Regmap.get r rs) args) m) eqn:EQ. - + exploit eval_cond_correct; eauto. intros. eapply RTL.max_reg_function_use. apply INSTR. auto. - intros. econstructor. econstructor. eassumption. unfold boolToValue, Values.Val.of_optbool. - destruct b; constructor; auto. - + eapply eval_cond_correct'; eauto. intros. eapply RTL.max_reg_function_use. apply INSTR. auto. - - monadInv H1. - destruct (Op.eval_condition c (map (fun r1 : positive => Registers.Regmap.get r1 rs) l0) m) eqn:EQN; - simplify. destruct b eqn:B. - + exploit eval_cond_correct; eauto. intros. eapply RTL.max_reg_function_use. apply INSTR. - simplify; tauto. intros. - econstructor. econstructor. eapply Verilog.erun_Vternary_true. eassumption. econstructor. auto. - auto. unfold Values.Val.normalize. - destruct (Registers.Regmap.get r rs) eqn:EQN2; constructor. - * assert (HPle1 : Ple r (RTL.max_reg_function f)) by (eapply RTL.max_reg_function_use; eauto; simpl; auto). - apply H in HPle1. inv HPle1. unfold valueToInt in H1. rewrite EQN2 in H1. inv H1. auto. - rewrite EQN2 in H1. discriminate. rewrite EQN2 in H2. discriminate. - * assert (HPle1 : Ple r (RTL.max_reg_function f)) by (eapply RTL.max_reg_function_use; eauto; simpl; auto). - apply H in HPle1. inv HPle1. rewrite EQN2 in H1. inv H1. rewrite EQN2 in H1. inv H1. auto. - rewrite EQN2 in H2. discriminate. - + exploit eval_cond_correct; eauto. intros. eapply RTL.max_reg_function_use. apply INSTR. - simplify; tauto. intros. - econstructor. econstructor. eapply Verilog.erun_Vternary_false. eassumption. econstructor. auto. - auto. unfold Values.Val.normalize. - destruct (Registers.Regmap.get r0 rs) eqn:EQN2; constructor. - * assert (HPle1 : Ple r0 (RTL.max_reg_function f)) by (eapply RTL.max_reg_function_use; eauto; simpl; auto). - apply H in HPle1. inv HPle1. unfold valueToInt in H1. rewrite EQN2 in H1. inv H1. auto. - rewrite EQN2 in H1. discriminate. rewrite EQN2 in H2. discriminate. - * assert (HPle1 : Ple r0 (RTL.max_reg_function f)) by (eapply RTL.max_reg_function_use; eauto; simpl; auto). - apply H in HPle1. inv HPle1. rewrite EQN2 in H1. inv H1. rewrite EQN2 in H1. inv H1. auto. - rewrite EQN2 in H2. discriminate. - + exploit eval_cond_correct'; eauto. intros. eapply RTL.max_reg_function_use. apply INSTR. - simplify; tauto. intros. inv H0. inv H1. destruct (Int.eq_dec x0 Int.zero). - econstructor. econstructor. eapply Verilog.erun_Vternary_false. - eassumption. econstructor. auto. subst. auto. constructor. - econstructor. econstructor. eapply Verilog.erun_Vternary_true. - eassumption. econstructor. auto. unfold valueToBool. pose proof n. apply Int.eq_false in n. - unfold uvalueToZ. unfold Int.eq in n. unfold zeq in *. - destruct (Int.unsigned x0 ==Z Int.unsigned Int.zero); try discriminate. - rewrite <- Z.eqb_neq in n0. rewrite Int.unsigned_zero in n0. rewrite n0. auto. - constructor. - Qed. - - (** The proof of semantic preservation for the translation of instructions - is a simulation argument based on diagrams of the following form: -<< - match_states - code st rs ---------------- State m st assoc - || | - || | - || | - \/ v - code st rs' --------------- State m st assoc' - match_states ->> - where [tr_code c data control fin rtrn st] is assumed to hold. - - The precondition and postcondition is that that should hold is [match_assocmaps rs assoc]. - *) - - Definition transl_instr_prop (instr : RTL.instruction) : Prop := - forall m asr asa fin rtrn st stmt trans res, - tr_instr fin rtrn st (m.(HTL.mod_stk)) instr stmt trans -> - exists asr' asa', - HTL.step tge (HTL.State res m st asr asa) Events.E0 (HTL.State res m st asr' asa'). - -Ltac name_goal name := refine ?[name]. - -Ltac unfold_merge := - unfold merge_assocmap; repeat (rewrite AssocMapExt.merge_add_assoc); - try (rewrite AssocMapExt.merge_base_1). - - Ltac tac0 := - match goal with - | [ |- HTL.exec_ram _ _ _ _ _ ] => constructor - | [ |- context[Verilog.merge_regs _ _] ] => unfold Verilog.merge_regs; cbn; unfold_merge - | [ |- context[reg_stack_based_pointers] ] => unfold reg_stack_based_pointers; intros - - | [ |- 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 - - | [ |- 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 _) ] => - apply regs_lessdef_add_greater; [> unfold Plt; lia | assumption] - - | [ 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[Registers.Regmap.get ?d (Registers.Regmap.set ?d _ _)] ] => - rewrite Registers.Regmap.gss - | [ |- context[Registers.Regmap.get ?s (Registers.Regmap.set ?d _ _)] ] => - let EQ := fresh "EQ" in - destruct (Pos.eq_dec s d) as [EQ|EQ]; - [> rewrite EQ | rewrite Registers.Regmap.gso; auto] - - | [ H : opt_val_value_lessdef _ _ |- _ ] => inv H - | [ H : context[Z.of_nat (Z.to_nat _)] |- _ ] => rewrite Z2Nat.id in H; [> solve crush |] - | [ H : _ ! _ = Some _ |- _] => setoid_rewrite H - | [ |- context[AssocMapExt.merge]] => progress unfold_merge - end. - - Ltac simplify_local := intros; unfold_constants; cbn in *; - repeat (nicify_hypotheses; nicify_goals; kill_bools; substpp); - cbn in *. - - Ltac simplify_val := repeat (simplify_local; unfold uvalueToZ, valueToPtr, Ptrofs.of_int, valueToInt, intToValue, - ptrToValue in *). - - Ltac crush_val := simplify_val; try discriminate; try congruence; try lia; liapp; try assumption. - - Ltac small_tac := repeat (crush_val; try array; try ptrofs); crush_val; auto. - Ltac big_tac := repeat (crush_val; try array; try ptrofs; try tac0); crush_val; auto. - - Lemma match_assocmaps_merge_empty: - forall f rs ars, - match_assocmaps f rs ars -> - match_assocmaps f rs (AssocMapExt.merge value empty_assocmap ars). - Proof. - inversion 1; subst; clear H. - constructor; intros. - rewrite merge_get_default2; auto. - Qed. - - Opaque AssocMap.get. - Opaque AssocMap.set. - Opaque AssocMapExt.merge. - Opaque Verilog.merge_arr. - - Lemma match_assocmaps_ext : - forall f rs ars1 ars2, - (forall x, Ple x (RTL.max_reg_function f) -> ars1 ! x = ars2 ! x) -> - match_assocmaps f rs ars1 -> - match_assocmaps f rs ars2. - Proof. - intros * YFRL YMATCH. - inv YMATCH. constructor; intros x' YPLE. - unfold "#", AssocMapExt.get_default in *. - rewrite <- YFRL by auto. - eauto. - Qed. - - 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. - inv CONST; assumption. - inv CONST; assumption. - (* processing of state *) - econstructor. - crush. - econstructor. - econstructor. - econstructor. - big_tac. - cbn. - - solve [inv MARR; big_tac]. - - (* inv MARR; big_tac. *) - inv MARR; big_tac; auto. - - - eapply match_assocmaps_ext; [|eauto]; intros. - repeat unfold_merge. rewrite AssocMap.gso by (unfold Ple in *; lia). - rewrite AssocMapExt.merge_base_1; auto. - - rewrite <- H1. unfold Verilog.merge_arrs. - rewrite !AssocMap.gcombine by auto. rewrite !AssocMap.gss. - setoid_rewrite H1. - repeat erewrite Verilog.merge_arr_empty2; eauto. - - inv CONST; cbn in *. constructor; cbn in *. - + repeat unfold_merge. rewrite AssocMap.gso by lia. - unfold_merge; auto. - + repeat unfold_merge. rewrite AssocMap.gso by lia. - unfold_merge; auto. - - Unshelve. exact tt. - Qed. - #[local] 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. - intros s f sp pc rs m op args res0 pc' v H H0 R1 MSTATE. - inv_state. inv MARR. - exploit eval_correct; eauto. intros. inversion H1. inversion H2. - econstructor. split. - apply Smallstep.plus_one. - eapply HTL.step_module; eauto. - inv CONST. assumption. - inv CONST. assumption. - econstructor; simpl; trivial. - constructor; trivial. - econstructor; simpl; eauto. - simpl. econstructor. econstructor. - apply H5. simplify. - - all: big_tac. - - - assert (HPle: Ple res0 (RTL.max_reg_function f)) - by (eapply RTL.max_reg_function_def; eauto; simpl; auto). - unfold Ple in HPle. lia. - - eapply match_assocmaps_merge_empty. eapply match_assocmaps_ext; intros. - unfold Ple in *. instantiate (1 := asr # res0 <- x). - destruct (peq res0 x1); subst. - + rewrite merge_get_default4 with (x := x); - apply AssocMap.gss. - + rewrite merge_get_default3; [now rewrite AssocMap.gso by auto|]. - rewrite AssocMap.gso by auto. - now rewrite AssocMap.gso by lia. - + now apply regs_lessdef_add_match. - - assert (HPle: Ple res0 (RTL.max_reg_function f)) - by (eapply RTL.max_reg_function_def; eauto; simpl; auto). - unfold Ple in HPle. lia. - - unfold Verilog.merge_arrs. - rewrite ! AssocMap.gcombine by auto. rewrite ! AssocMap.gss. - erewrite ! Verilog.merge_arr_empty2; eauto. - erewrite ! Verilog.merge_arr_empty2; eauto. - - assumption. - - eapply op_stack_based; eauto. - - assert (HPle: Ple res0 (RTL.max_reg_function f)) - by (eapply RTL.max_reg_function_def; eauto; simpl; auto). - unfold Ple in *. - inv CONST. constructor; cbn. - repeat rewrite merge_get_default3 by solve [auto | lia]. - rewrite merge_get_default3; [eauto | ]. - repeat rewrite AssocMap.gso by solve [auto | lia]. auto. - repeat rewrite merge_get_default3 by solve [auto | lia]. - rewrite merge_get_default3; [eauto | ]. - repeat rewrite AssocMap.gso by solve [auto | lia]. auto. - Unshelve. apply tt. - Qed. - #[local] 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 _ _ = _ |- _ ] => inv H - | [ _ : 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; crush; tac; crush - end. - - Lemma offset_expr_ok : - forall v z, (Z.to_nat - (Integers.Ptrofs.unsigned - (Integers.Ptrofs.divu - (Integers.Ptrofs.add (Integers.Ptrofs.repr (uvalueToZ v)) - (Integers.Ptrofs.of_int (Integers.Int.repr z))) - (Integers.Ptrofs.repr 4))) - = valueToNat (Int.divu (Int.add v (ZToValue z)) (ZToValue 4))). - Proof. - simplify_val. unfold valueToNat. unfold Int.divu, Ptrofs.divu. - pose proof Integers.Ptrofs.agree32_add as AGR. - unfold Integers.Ptrofs.agree32 in AGR. - assert (Ptrofs.unsigned (Ptrofs.add (Ptrofs.repr (Int.unsigned v)) - (Ptrofs.repr (Int.unsigned (Int.repr z)))) = - Int.unsigned (Int.add v (ZToValue z))). - apply AGR; auto. - apply Ptrofs.unsigned_repr. apply Int.unsigned_range_2. - apply Ptrofs.unsigned_repr. apply Int.unsigned_range_2. - rewrite H. replace (Ptrofs.unsigned (Ptrofs.repr 4)) with 4. - replace (Int.unsigned (ZToValue 4)) with 4. - pose proof Ptrofs.agree32_repr. unfold Ptrofs.agree32 in *. - rewrite H0. trivial. auto. - unfold ZToValue. symmetry. apply Int.unsigned_repr. - unfold_constants. lia. - unfold ZToValue. symmetry. apply Int.unsigned_repr. - unfold_constants. lia. - Qed. - - Lemma offset_expr_ok_2 : - forall v0 v1 z0 z1, - (Z.to_nat - (Integers.Ptrofs.unsigned - (Integers.Ptrofs.divu - (Integers.Ptrofs.add (Integers.Ptrofs.repr (uvalueToZ v0)) - (Integers.Ptrofs.of_int - (Integers.Int.add - (Integers.Int.mul (valueToInt v1) (Integers.Int.repr z1)) - (Integers.Int.repr z0)))) (Ptrofs.repr 4)))) - = valueToNat (Int.divu (Int.add (Int.add v0 (ZToValue z0)) - (Int.mul v1 (ZToValue z1))) (ZToValue 4)). - intros. unfold ZToValue, valueToNat, valueToInt, Ptrofs.divu, Int.divu, Ptrofs.of_int. - - assert (H : (Ptrofs.unsigned - (Ptrofs.add (Ptrofs.repr (uvalueToZ v0)) - (Ptrofs.of_int (Int.add (Int.mul (valueToInt v1) (Int.repr z1)) (Int.repr z0)))) / - Ptrofs.unsigned (Ptrofs.repr 4)) - = (Int.unsigned (Int.add (Int.add v0 (Int.repr z0)) (Int.mul v1 (Int.repr z1))) / - Int.unsigned (Int.repr 4))). - { unfold ZToValue, valueToNat, valueToInt, Ptrofs.divu, Int.divu, Ptrofs.of_int. - rewrite Ptrofs.unsigned_repr by (unfold_constants; lia). - rewrite Int.unsigned_repr by (unfold_constants; lia). - - unfold Ptrofs.of_int. rewrite Int.add_commut. - pose proof Integers.Ptrofs.agree32_add as AGR. unfold Ptrofs.agree32 in *. - erewrite AGR. - 3: { unfold uvalueToZ. rewrite Ptrofs.unsigned_repr. trivial. apply Int.unsigned_range_2. } - 3: { rewrite Ptrofs.unsigned_repr. trivial. apply Int.unsigned_range_2. } - rewrite Int.add_assoc. trivial. auto. - } - - rewrite <- H. auto. - - Qed. - - Lemma offset_expr_ok_3 : - forall OFFSET, - Z.to_nat (Ptrofs.unsigned (Ptrofs.divu OFFSET (Ptrofs.repr 4))) - = valueToNat (ZToValue (Ptrofs.unsigned OFFSET / 4)). - Proof. auto. Qed. - - 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. inv_arr_access. *) - - (* + (** Preamble *) *) - (* inv MARR. inv CONST. crush. *) - - (* unfold Op.eval_addressing in H0. *) - (* 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; crush. *) - - (* rewrite ARCHI in H1. crush. *) - (* subst. *) - - (* pose proof MASSOC as MASSOC'. *) - (* inv MASSOC'. *) - (* pose proof (H0 r0). *) - (* assert (HPler0 : Ple r0 (RTL.max_reg_function f)) *) - (* by (eapply RTL.max_reg_function_use; eauto; crush; eauto). *) - (* apply H0 in HPler0. *) - (* inv HPler0; try congruence. *) - (* rewrite EQr0 in H11. *) - (* inv H11. *) - - (* unfold check_address_parameter_signed in *; *) - (* unfold check_address_parameter_unsigned in *; crush. *) - - (* remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (uvalueToZ 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. *) - (* { apply Mem.load_valid_access in H1. unfold Mem.valid_access in *. simplify. *) - (* apply Zdivide_mod. assumption. } *) - - (* (** Read bounds proof *) *) - (* assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH. *) - (* { destruct (Integers.Ptrofs.unsigned OFFSET - 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. inv_arr_access. *) - - (* + (** Preamble *) *) - (* inv MARR. inv CONST. crush. *) - - (* unfold Op.eval_addressing in H0. *) - (* 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; crush. *) - - (* rewrite ARCHI in H1. crush. *) - (* subst. *) - - (* pose proof MASSOC as MASSOC'. *) - (* inv MASSOC'. *) - (* pose proof (H0 r0). *) - (* assert (HPler0 : Ple r0 (RTL.max_reg_function f)) *) - (* by (eapply RTL.max_reg_function_use; eauto; crush; eauto). *) - (* apply H8 in HPler0. *) - (* inv HPler0; try congruence. *) - (* rewrite EQr0 in H11. *) - (* inv H11. *) - (* clear H0. clear H8. *) - - (* unfold check_address_parameter_unsigned in *; *) - (* unfold check_address_parameter_signed in *; crush. *) - - (* remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (uvalueToZ 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. *) - (* { apply Mem.store_valid_access_3 in H1. unfold Mem.valid_access in *. simplify. *) - (* apply Zdivide_mod. 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 H11. *) - (* exact H11. 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 H11. inv H11. congruence. *) - - (* constructor; simplify. unfold Verilog.merge_regs. unfold_merge. *) - (* rewrite AssocMap.gso. *) - (* assumption. lia. *) - (* unfold Verilog.merge_regs. unfold_merge. *) - (* rewrite AssocMap.gso. *) - (* assumption. lia. *) - - (* + (** Preamble *) *) - (* inv MARR. inv CONST. crush. *) - - (* unfold Op.eval_addressing in H0. *) - (* 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; crush. *) - - (* rewrite ARCHI in H1. crush. *) - (* subst. *) - (* clear RSBPr1. *) - - (* pose proof MASSOC as MASSOC'. *) - (* inv 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; crush; 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 H11 in HPler1. *) - (* inv HPler0; inv HPler1; try congruence. *) - (* rewrite EQr0 in H13. *) - (* rewrite EQr1 in H14. *) - (* inv H13. inv H14. *) - (* clear H0. clear H8. clear H11. *) - - (* unfold check_address_parameter_signed in *; *) - (* unfold check_address_parameter_unsigned in *; crush. *) - - (* remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (uvalueToZ 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. *) - (* { apply Mem.store_valid_access_3 in H1. unfold Mem.valid_access in *. simplify. *) - (* apply Zdivide_mod. 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 H14. *) - (* exact H14. 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 H14. inv H14. congruence. *) - - (* constructor; simplify. unfold Verilog.merge_regs. unfold_merge. rewrite AssocMap.gso. *) - (* assumption. lia. *) - (* unfold Verilog.merge_regs. unfold_merge. rewrite AssocMap.gso. *) - (* assumption. lia. *) - - (* + inv MARR. inv CONST. crush. *) - - (* unfold Op.eval_addressing in H0. *) - (* destruct (Archi.ptr64) eqn:ARCHI; crush. *) - (* rewrite ARCHI in H0. crush. *) - - (* unfold check_address_parameter_unsigned in *; *) - (* 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. *) - (* rewrite Integers.Ptrofs.add_zero_l in H1. *) - - (* remember i0 as OFFSET. *) - - (* (** Modular preservation proof *) *) - (* assert (Integers.Ptrofs.unsigned OFFSET mod 4 = 0) as MOD_PRESERVE. *) - (* { apply Mem.store_valid_access_3 in H1. unfold Mem.valid_access in *. simplify. *) - (* apply Zdivide_mod. 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. inv H0. congruence. *) - - (* constructor; simplify. unfold Verilog.merge_regs. unfold_merge. rewrite AssocMap.gso. *) - (* assumption. lia. *) - (* unfold Verilog.merge_regs. unfold_merge. rewrite AssocMap.gso. *) - (* assumption. lia. *) - - (* Unshelve. *) - (* exact tt. *) - (* exact (Values.Vint (Int.repr 0)). *) - (* exact tt. *) - (* exact (Values.Vint (Int.repr 0)). *) - (* exact tt. *) - (* exact (Values.Vint (Int.repr 0)). *) - (* Qed. *) Admitted. - #[local] 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. - destruct b. - - eexists. split. apply Smallstep.plus_one. - clear H33. - eapply HTL.step_module; eauto. - inv CONST; assumption. - inv CONST; assumption. - econstructor; simpl; trivial. - constructor; trivial. - eapply Verilog.erun_Vternary_true; simpl; eauto. - eapply eval_cond_correct; eauto. intros. - intros. eapply RTL.max_reg_function_use. apply H22. auto. - econstructor. auto. - simpl. econstructor. constructor. unfold Verilog.merge_regs. unfold_merge. simpl. - unfold_merge. apply AssocMap.gss. - - inv MARR. inv CONST. - big_tac. - (* constructor; rewrite AssocMap.gso; simplify; try assumption; lia. *) - (* - eexists. split. apply Smallstep.plus_one. *) - (* clear H32. *) - (* eapply HTL.step_module; eauto. *) - (* inv CONST; assumption. *) - (* inv CONST; assumption. *) - (* econstructor; simpl; trivial. *) - (* constructor; trivial. *) - (* eapply Verilog.erun_Vternary_false; simpl; eauto. *) - (* eapply eval_cond_correct; eauto. intros. *) - (* intros. eapply RTL.max_reg_function_use. apply H22. auto. *) - (* econstructor. auto. *) - (* simpl. econstructor. constructor. unfold Verilog.merge_regs. unfold_merge. simpl. *) - (* apply AssocMap.gss. *) - - (* inv MARR. inv CONST. *) - (* big_tac. *) - (* constructor; rewrite AssocMap.gso; simplify; try assumption; lia. *) - - (* Unshelve. all: exact tt. *) - (* Qed. *) Admitted. - #[local] 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. *) - - (* #[local] 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. - inv CONST; assumption. - inv CONST; assumption. - constructor. - econstructor; simpl; trivial. - econstructor; simpl; trivial. - constructor. - econstructor; simpl; trivial. - constructor. - - constructor. constructor. constructor. - - unfold state_st_wf in WF; big_tac; eauto. - destruct wf1 as [HCTRL HDATA]. apply HCTRL. - apply AssocMapExt.elements_iff. eexists. - match goal with H: control ! pc = Some _ |- _ => apply H end. - - apply HTL.step_finish. - unfold Verilog.merge_regs. - unfold_merge; simpl. - (* rewrite AssocMap.gso. *) - (* apply AssocMap.gss. lia. *) - (* apply AssocMap.gss. *) - (* rewrite Events.E0_left. reflexivity. *) - - (* constructor; auto. *) - (* constructor. *) - - (* (* FIXME: Duplication *) *) - (* - econstructor. split. *) - (* eapply Smallstep.plus_two. *) - (* eapply HTL.step_module; eauto. *) - (* inv CONST; assumption. *) - (* inv CONST; assumption. *) - (* constructor. *) - (* econstructor; simpl; trivial. *) - (* econstructor; simpl; trivial. *) - (* constructor. constructor. constructor. *) - (* constructor. constructor. constructor. *) - (* constructor. *) - - (* unfold state_st_wf in WF; big_tac; eauto. *) - - (* destruct wf1 as [HCTRL HDATA]. apply HCTRL. *) - (* apply AssocMapExt.elements_iff. eexists. *) - (* match goal with H: control ! pc = Some _ |- _ => apply H end. *) - - (* apply HTL.step_finish. *) - (* unfold Verilog.merge_regs. *) - (* unfold_merge. *) - (* unfold_merge. *) - (* rewrite AssocMap.gso. *) - (* apply AssocMap.gss. simpl; lia. *) - (* apply AssocMap.gss. *) - (* rewrite Events.E0_left. trivial. *) - - (* constructor; auto. *) - - (* simpl. inversion MASSOC. subst. *) - (* unfold find_assocmap, AssocMapExt.get_default. rewrite AssocMap.gso. *) - (* apply H1. eapply RTL.max_reg_function_use. eauto. simpl; tauto. *) - (* assert (HPle : Ple r (RTL.max_reg_function f)). *) - (* eapply RTL.max_reg_function_use. eassumption. simpl; auto. *) - (* apply ZExtra.Ple_not_eq. apply ZExtra.Ple_Plt_Succ. assumption. *) - - (* Unshelve. *) - (* all: constructor. *) - (* Qed. *) Admitted. - #[local] 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. - - inversion MSTATE; subst. inversion TF; subst. - econstructor. split. apply Smallstep.plus_one. - eapply HTL.step_call. crush. - - apply match_state with (sp' := stk); eauto. - - all: big_tac. - - apply regs_lessdef_add_greater. unfold Plt; lia. - apply regs_lessdef_add_greater. unfold Plt; lia. - apply regs_lessdef_add_greater. unfold Plt; lia. - apply init_reg_assoc_empty. - - constructor. - - 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. } - ptrofs. rewrite LOAD. - rewrite ALLOC. - repeat constructor. - - ptrofs. rewrite LOAD. - repeat constructor. - - unfold reg_stack_based_pointers. intros. - unfold RTL.init_regs; crush. - destruct (RTL.fn_params f); - rewrite Registers.Regmap.gi; constructor. - - unfold arr_stack_based_pointers. intros. - crush. - 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 Mem.alloc in H. - inv H. - crush. - unfold Mem.load. - intros. - match goal with | |- context[if ?x then _ else _] => destruct x end; try congruence. - inv v0. unfold Mem.range_perm in H4. - unfold Mem.perm in H4. crush. - unfold Mem.perm_order' in H4. - small_tac. - exploit (H4 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. - crush. - apply proj_sumbool_true in H10. lia. - - unfold Mem.alloc in H. - inv H. - crush. - unfold Mem.store. - intros. - match goal with | |- context[if ?x then _ else _] => destruct x end; try congruence. - inv v0. unfold Mem.range_perm in H4. - unfold Mem.perm in H4. crush. - unfold Mem.perm_order' in H4. - small_tac. - exploit (H4 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. - crush. - apply proj_sumbool_true in H10. lia. - constructor. simplify. rewrite AssocMap.gss. - simplify. rewrite AssocMap.gso. apply AssocMap.gss. simplify. lia. - Opaque Mem.alloc. - Opaque Mem.load. - Opaque Mem.store. - Qed. - #[local] 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. - #[local] Hint Resolve transl_returnstate_correct : htlproof. - - Lemma option_inv : - forall A x y, - @Some A x = Some y -> x = y. - 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. - - Lemma transl_initial_states : - forall s1 : Smallstep.state (RTL.semantics prog), - Smallstep.initial_state (RTL.semantics prog) s1 -> - exists s2 : Smallstep.state (HTL.semantics tprog), - 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. - 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 H6. - - constructor. - 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. inv H7. assumption. - Qed. - #[local] Hint Resolve transl_initial_states : htlproof. - - Lemma transl_final_states : - forall (s1 : Smallstep.state (RTL.semantics prog)) - (s2 : Smallstep.state (HTL.semantics tprog)) - (r : Integers.Int.int), - match_states s1 s2 -> - 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 MF. constructor. reflexivity. - Qed. - #[local] Hint Resolve transl_final_states : htlproof. - - 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. - #[local] 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. diff --git a/src/hls/IfConversion.v b/src/hls/IfConversion.v index 8ca9cb6..ecafc13 100644 --- a/src/hls/IfConversion.v +++ b/src/hls/IfConversion.v @@ -247,7 +247,7 @@ Definition ifconv_list (headers: list node) (c: code) := Definition if_convert_code (c: code) iflist := fold_left (fun s n => if_convert c s (fst n) (snd n)) iflist c. -Definition transf_function (l: option if_conv_t) (f: function) : function := +Definition transf_function (f: function) : function := let (b, _) := build_bourdoncle f in let b' := get_loops b in let iflist := ifconv_list b' f.(fn_code) in @@ -274,8 +274,8 @@ Section TRANSF_PROGRAM. End TRANSF_PROGRAM. -Definition transf_fundef (l: option if_conv_t) (fd: fundef) : fundef := - transf_fundef (transf_function l) fd. +Definition transf_fundef (fd: fundef) : fundef := + transf_fundef transf_function fd. -Definition transf_program (l: PTree.t if_conv_t) (p: program) : program := - transform_program_data transf_fundef l p. +Definition transf_program (p: program) : program := + transform_program transf_fundef p. diff --git a/src/hls/IfConversionproof.v b/src/hls/IfConversionproof.v index 3c9d2bd..eedb318 100644 --- a/src/hls/IfConversionproof.v +++ b/src/hls/IfConversionproof.v @@ -50,8 +50,8 @@ Require Import vericert.hls.Predicate. Variant match_stackframe : stackframe -> stackframe -> Prop := | match_stackframe_init : - forall res f tf sp pc rs p l - (TF: transf_function l f = tf), + forall res f tf sp pc rs p + (TF: transf_function f = tf), match_stackframe (Stackframe res f sp pc rs p) (Stackframe res tf sp pc rs p). Definition bool_order (b: bool): nat := if b then 1 else 0. @@ -234,8 +234,8 @@ Proof. Qed. Lemma transf_spec_correct : - forall f pc l, - if_conv_spec f.(fn_code) (transf_function l f).(fn_code) pc. + forall f pc, + if_conv_spec f.(fn_code) (transf_function f).(fn_code) pc. Proof. intros; unfold transf_function; destruct_match; cbn. unfold if_convert_code. @@ -309,8 +309,8 @@ Section CORRECTNESS. Variant match_states : option SeqBB.t -> state -> state -> Prop := | match_state_some : - forall stk stk' f tf sp pc rs p m b pc0 rs0 p0 m0 l - (TF: transf_function l f = tf) + forall stk stk' f tf sp pc rs p m b pc0 rs0 p0 m0 + (TF: transf_function f = tf) (STK: Forall2 match_stackframe stk stk') (* This can be improved with a recursive relation for a more general structure of the if-conversion proof. *) @@ -322,13 +322,13 @@ Section CORRECTNESS. (SIM: step ge (State stk f sp pc0 rs0 p0 m0) E0 (State stk f sp pc rs p m)), match_states (Some b) (State stk f sp pc rs p m) (State stk' tf sp pc0 rs0 p0 m0) | match_state_none : - forall stk stk' f tf sp pc rs p m l - (TF: transf_function l f = tf) + forall stk stk' f tf sp pc rs p m + (TF: transf_function f = tf) (STK: Forall2 match_stackframe stk stk'), match_states None (State stk f sp pc rs p m) (State stk' tf sp pc rs p m) | match_callstate : - forall cs cs' f tf args m l - (TF: transf_fundef l f = tf) + forall cs cs' f tf args m + (TF: transf_fundef f = tf) (STK: Forall2 match_stackframe cs cs'), match_states None (Callstate cs f args m) (Callstate cs' tf args m) | match_returnstate : @@ -337,7 +337,13 @@ Section CORRECTNESS. match_states None (Returnstate cs v m) (Returnstate cs' v m). Definition match_prog (p: program) (tp: program) := - match_program (fun _ f tf => exists l, transf_fundef l f = tf) eq p tp. + match_program (fun _ f tf => tf = transf_fundef f) eq p tp. + + Lemma transf_program_match: + forall prog, match_prog prog (transf_program prog). + Proof. + intros. eapply Linking.match_transform_program_contextual. auto. + Qed. Context (TRANSL : match_prog prog tprog). @@ -354,34 +360,34 @@ Section CORRECTNESS. Qed. Lemma function_ptr_translated: - forall b f l, + forall b f, Genv.find_funct_ptr ge b = Some f -> - Genv.find_funct_ptr tge b = Some (transf_fundef l f). + Genv.find_funct_ptr tge b = Some (transf_fundef f). Proof. intros. exploit (Genv.find_funct_ptr_match TRANSL); eauto. crush. Qed. Lemma sig_transf_function: - forall (f tf: fundef) l, - funsig (transf_fundef l f) = funsig f. + forall (f tf: fundef), + funsig (transf_fundef f) = funsig f. Proof using. unfold transf_fundef. unfold AST.transf_fundef; intros. destruct f. unfold transf_function. destruct_match. auto. auto. Qed. Lemma functions_translated: - forall (v: Values.val) (f: GibleSeq.fundef) l, + forall (v: Values.val) (f: GibleSeq.fundef), Genv.find_funct ge v = Some f -> - Genv.find_funct tge v = Some (transf_fundef l f). + Genv.find_funct tge v = Some (transf_fundef f). Proof using TRANSL. intros. exploit (Genv.find_funct_match TRANSL); eauto. simplify. eauto. Qed. Lemma find_function_translated: - forall ros rs f l, + forall ros rs f, find_function ge ros rs = Some f -> - find_function tge ros rs = Some (transf_fundef l f). + find_function tge ros rs = Some (transf_fundef f). Proof using TRANSL. Ltac ffts := match goal with | [ H: forall _, Val.lessdef _ _, r: Registers.reg |- _ ] => @@ -410,7 +416,7 @@ Section CORRECTNESS. replace (prog_main tprog) with (prog_main prog). rewrite symbols_preserved; eauto. symmetry; eapply Linking.match_program_main; eauto. eauto. erewrite sig_transf_function; eauto. econstructor. auto. auto. - Unshelve. exact None. + Unshelve. Qed. Lemma transf_final_states : @@ -507,8 +513,8 @@ Section CORRECTNESS. Qed. Lemma fn_all_eq : - forall f tf l, - transf_function l f = tf -> + forall f tf, + transf_function f = tf -> fn_stacksize f = fn_stacksize tf /\ fn_sig f = fn_sig tf /\ fn_params f = fn_params tf @@ -521,16 +527,16 @@ Section CORRECTNESS. Ltac func_info := match goal with - H: transf_function _ _ = _ |- _ => + H: transf_function _ = _ |- _ => let H2 := fresh "ALL_EQ" in - pose proof (fn_all_eq _ _ _ H) as H2; simplify + pose proof (fn_all_eq _ _ H) as H2; simplify end. Lemma step_cf_eq : - forall stk stk' f tf sp pc rs pr m cf s t pc' l, + forall stk stk' f tf sp pc rs pr m cf s t pc', step_cf_instr ge (State stk f sp pc rs pr m) cf t s -> Forall2 match_stackframe stk stk' -> - transf_function l f = tf -> + transf_function f = tf -> exists s', step_cf_instr tge (State stk' tf sp pc' rs pr m) cf t s' /\ match_states None s s'. Proof. @@ -778,19 +784,19 @@ Section CORRECTNESS. Qed. Lemma match_none_correct : - forall t s1' stk f sp pc rs m pr rs' m' bb pr' cf stk' l, + forall t s1' stk f sp pc rs m pr rs' m' bb pr' cf stk', (fn_code f) ! pc = Some bb -> SeqBB.step ge sp (Iexec (mki rs pr m)) bb (Iterm (mki rs' pr' m') cf) -> step_cf_instr ge (State stk f sp pc rs' pr' m') cf t s1' -> Forall2 match_stackframe stk stk' -> exists b' s2', - (plus step tge (State stk' (transf_function l f) sp pc rs pr m) t s2' \/ - star step tge (State stk' (transf_function l f) sp pc rs pr m) t s2' + (plus step tge (State stk' (transf_function f) sp pc rs pr m) t s2' \/ + star step tge (State stk' (transf_function f) sp pc rs pr m) t s2' /\ ltof (option SeqBB.t) measure b' None) /\ match_states b' s1' s2'. Proof. intros * H H0 H1 STK. - pose proof (transf_spec_correct f pc l) as X; inv X. + pose proof (transf_spec_correct f pc) as X; inv X. - apply sim_plus. rewrite H in H2. symmetry in H2. exploit step_cf_eq; eauto; simplify. do 3 econstructor. apply plus_one. econstructor; eauto. @@ -800,7 +806,7 @@ Section CORRECTNESS. destruct (cf_wf_dec x b' cf pc'); subst; simplify. + inv H1. exploit exec_if_conv; eauto; simplify. - apply sim_star. exists (Some b'). exists (State stk' (transf_function l f) sp pc rs pr m). + apply sim_star. exists (Some b'). exists (State stk' (transf_function f) sp pc rs pr m). simplify; try (unfold ltof; auto). apply star_refl. econstructor; auto. simplify. econstructor; eauto. @@ -820,21 +826,21 @@ Section CORRECTNESS. Qed. Lemma match_some_correct: - forall t s1' s f sp pc rs m pr rs' m' bb pr' cf stk' b0 pc1 rs1 p0 m1 l, + forall t s1' s f sp pc rs m pr rs' m' bb pr' cf stk' b0 pc1 rs1 p0 m1, step ge (State s f sp pc rs pr m) t s1' -> (fn_code f) ! pc = Some bb -> SeqBB.step ge sp (Iexec (mki rs pr m)) bb (Iterm (mki rs' pr' m') cf) -> step_cf_instr ge (State s f sp pc rs' pr' m') cf t s1' -> Forall2 match_stackframe s stk' -> (fn_code f) ! pc = Some b0 -> - sem_extrap (transf_function l f) pc1 sp (Iexec (mki rs pr m)) (Iexec (mki rs1 p0 m1)) b0 -> + sem_extrap (transf_function f) pc1 sp (Iexec (mki rs pr m)) (Iexec (mki rs1 p0 m1)) b0 -> (forall b', f.(fn_code)!pc1 = Some b' -> - exists tb, (transf_function l f).(fn_code)!pc1 = Some tb /\ if_conv_replace pc b0 b' tb) -> + exists tb, (transf_function f).(fn_code)!pc1 = Some tb /\ if_conv_replace pc b0 b' tb) -> step ge (State s f sp pc1 rs1 p0 m1) E0 (State s f sp pc rs pr m) -> exists b' s2', - (plus step tge (State stk' (transf_function l f) sp pc1 rs1 p0 m1) t s2' \/ - star step tge (State stk' (transf_function l f) sp pc1 rs1 p0 m1) t s2' /\ + (plus step tge (State stk' (transf_function f) sp pc1 rs1 p0 m1) t s2' \/ + star step tge (State stk' (transf_function f) sp pc1 rs1 p0 m1) t s2' /\ ltof (option SeqBB.t) measure b' (Some b0)) /\ match_states b' s1' s2'. Proof. intros * H H0 H1 H2 STK IS_B EXTRAP IS_TB SIM. @@ -863,14 +869,14 @@ Section CORRECTNESS. match goal with H: context[match_states] |- _ => inv H end. - eauto using match_some_correct. - eauto using match_none_correct. - - apply sim_plus. remember (transf_function l f) as tf. symmetry in Heqtf. func_info. + - apply sim_plus. remember (transf_function f) as tf. symmetry in Heqtf. func_info. exists (@None (list instr)). eexists. split. apply plus_one. constructor; eauto. rewrite <- H1. eassumption. rewrite <- H4. rewrite <- H2. econstructor; auto. - apply sim_plus. exists (@None (list instr)). eexists. split. apply plus_one. constructor; eauto. constructor; auto. - - apply sim_plus. remember (transf_function None f) as tf. symmetry in Heqtf. func_info. + - apply sim_plus. remember (transf_function f) as tf. symmetry in Heqtf. func_info. exists (@None (list instr)). inv STK. inv H7. eexists. split. apply plus_one. constructor. econstructor; auto. Qed. diff --git a/src/hls/Memorygen.v b/src/hls/Memorygen.v deleted file mode 100644 index 34e264d..0000000 --- a/src/hls/Memorygen.v +++ /dev/null @@ -1,3201 +0,0 @@ -(* - * Vericert: Verified high-level synthesis. - * Copyright (C) 2021 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 . - *) - -Require Import Coq.micromega.Lia. - -Require Import compcert.backend.Registers. -Require Import compcert.common.AST. -Require Import compcert.common.Globalenvs. -Require compcert.common.Memory. -Require Import compcert.common.Values. -Require Import compcert.lib.Floats. -Require Import compcert.lib.Integers. -Require Import compcert.lib.Maps. -Require compcert.common.Smallstep. -Require compcert.verilog.Op. - -Require Import vericert.common.Vericertlib. -Require Import vericert.hls.ValueInt. -Require Import vericert.hls.Verilog. -Require Import vericert.hls.HTL. -Require Import vericert.hls.AssocMap. -Require Import vericert.hls.Array. - -Local Open Scope positive. -Local Open Scope assocmap. - -#[local] Hint Resolve max_reg_stmnt_le_stmnt_tree : mgen. -#[local] Hint Resolve max_reg_stmnt_lt_stmnt_tree : mgen. -#[local] Hint Resolve max_stmnt_lt_module : mgen. - -Lemma int_eq_not_false : forall x, Int.eq x (Int.not x) = false. -Proof. - intros. unfold Int.eq. - rewrite Int.unsigned_not. - replace Int.max_unsigned with 4294967295%Z by crush. - assert (X: forall x, (x <> 4294967295 - x)%Z) by lia. - specialize (X (Int.unsigned x)). apply zeq_false; auto. -Qed. - -Definition transf_maps state ram in_ (dc: PTree.t stmnt * PTree.t stmnt) := - match dc, in_ with - | (d, c), (i, n) => - match PTree.get i d, PTree.get i c with - | Some (Vnonblock (Vvari r e1) e2), Some c_s => - if r =? (ram_mem ram) then - let nd := Vseq (Vnonblock (Vvar (ram_u_en ram)) (Vunop Vnot (Vvar (ram_u_en ram)))) - (Vseq (Vnonblock (Vvar (ram_wr_en ram)) (Vlit (ZToValue 1))) - (Vseq (Vnonblock (Vvar (ram_d_in ram)) e2) - (Vnonblock (Vvar (ram_addr ram)) e1))) - in - (PTree.set i nd d, c) - else dc - | Some (Vnonblock (Vvar e1) (Vvari r e2)), Some (Vnonblock (Vvar st') e3) => - if (r =? (ram_mem ram)) && (st' =? state) && (Z.pos n <=? Int.max_unsigned)%Z - && (e1 dc - end - end. - -Lemma transf_maps_wf : - forall state ram d c d' c' i, - map_well_formed c /\ map_well_formed d -> - transf_maps state ram i (d, c) = (d', c') -> - map_well_formed c' /\ map_well_formed d'. -Proof. - unfold transf_maps; intros; repeat destruct_match; - match goal with - | H: (_, _) = (_, _) |- _ => inv H - end; auto. - unfold map_well_formed. - simplify; intros. - destruct (Pos.eq_dec p0 p1); subst; auto. - destruct (Pos.eq_dec p p1); subst. unfold map_well_formed in *. - apply AssocMap.elements_correct in Heqo. - eapply in_map with (f := fst) in Heqo. simplify. - apply H1 in Heqo. auto. - apply AssocMapExt.elements_iff in H3. inv H3. - repeat rewrite AssocMap.gso in H8 by lia. - apply AssocMap.elements_correct in H8. - eapply in_map with (f := fst) in H8. simplify. - unfold map_well_formed in *. apply H0 in H8. auto. - apply AssocMapExt.elements_iff in H3. inv H3. - destruct (Pos.eq_dec p0 p1); subst; auto. - destruct (Pos.eq_dec p p1); subst. unfold map_well_formed in *. - apply AssocMap.elements_correct in Heqo. - eapply in_map with (f := fst) in Heqo. simplify. - apply H1 in Heqo. auto. - repeat rewrite AssocMap.gso in H8 by lia. - apply AssocMap.elements_correct in H8. - eapply in_map with (f := fst) in H8. simplify. - unfold map_well_formed in *. apply H1 in H8. auto. - unfold map_well_formed in *; simplify; intros. - destruct (Pos.eq_dec p0 p1); subst; auto. - destruct (Pos.eq_dec p p1); subst. unfold map_well_formed in *. - apply AssocMap.elements_correct in Heqo. - eapply in_map with (f := fst) in Heqo. simplify. - apply H1 in Heqo. auto. - apply AssocMapExt.elements_iff in H. inv H. - repeat rewrite AssocMap.gso in H2 by lia. - apply AssocMap.elements_correct in H2. - eapply in_map with (f := fst) in H2. simplify. - unfold map_well_formed in *. apply H1 in H2. auto. -Qed. - -Definition max_pc {A: Type} (m: PTree.t A) := - fold_right Pos.max 1%positive (map fst (PTree.elements m)). - -Fixpoint zip_range {A: Type} n (l: list A) {struct l} := - match l with - | nil => nil - | a :: b => (a, n) :: zip_range (n+1) b - end. - -Lemma zip_range_fst_idem : forall A (l: list A) a, map fst (zip_range a l) = l. -Proof. induction l; crush. Qed. - -Lemma zip_range_not_in_snd : - forall A (l: list A) a n, a < n -> ~ In a (map snd (zip_range n l)). -Proof. - unfold not; induction l; crush. - inv H0; try lia. eapply IHl. - assert (X: a0 < n + 1) by lia. apply X; auto. auto. -Qed. - -Lemma zip_range_snd_no_repet : - forall A (l: list A) a, list_norepet (map snd (zip_range a l)). -Proof. - induction l; crush; constructor; auto; []. - apply zip_range_not_in_snd; lia. -Qed. - -Lemma zip_range_in : - forall A (l: list A) a n i, In (a, i) (zip_range n l) -> In a l. -Proof. - induction l; crush. inv H. inv H0. auto. right. eapply IHl; eauto. -Qed. - -Lemma zip_range_not_in : - forall A (l: list A) a i n, ~ In a l -> ~ In (a, i) (zip_range n l). -Proof. - unfold not; induction l; crush. inv H0. inv H1. apply H. left. auto. - apply H. right. eapply zip_range_in; eauto. -Qed. - -Lemma zip_range_no_repet : - forall A (l: list A) a, list_norepet l -> list_norepet (zip_range a l). -Proof. - induction l; simplify; constructor; - match goal with H: list_norepet _ |- _ => inv H end; - [apply zip_range_not_in|]; auto. -Qed. - -Definition transf_code state ram d c := - fold_right (transf_maps state ram) (d, c) - (zip_range (Pos.max (max_pc c) (max_pc d) + 1) - (map fst (PTree.elements d))). - -Lemma transf_code_wf' : - forall l c d state ram c' d', - map_well_formed c /\ map_well_formed d -> - fold_right (transf_maps state ram) (d, c) l = (d', c') -> - map_well_formed c' /\ map_well_formed d'. -Proof. - induction l; intros; simpl in *. inv H0; auto. - remember (fold_right (transf_maps state ram) (d, c) l). destruct p. - apply transf_maps_wf in H0. auto. eapply IHl; eauto. -Qed. - -Lemma transf_code_wf : - forall c d state ram c' d', - map_well_formed c /\ map_well_formed d -> - transf_code state ram d c = (d', c') -> - map_well_formed c' /\ map_well_formed d'. -Proof. eauto using transf_code_wf'. Qed. - -Lemma ram_wf : - forall x, - x + 1 < x + 2 /\ x + 2 < x + 3 /\ x + 3 < x + 4 /\ x + 4 < x + 5 /\ x + 5 < x + 6. -Proof. lia. Qed. - -Lemma module_ram_wf' : - forall m addr, - addr = max_reg_module m + 1 -> - mod_clk m < addr. -Proof. unfold max_reg_module; lia. Qed. - -Definition transf_module (m: module): module. - refine ( - let max_reg := max_reg_module m in - let addr := max_reg+1 in - let en := max_reg+2 in - let d_in := max_reg+3 in - let d_out := max_reg+4 in - let wr_en := max_reg+5 in - let u_en := max_reg+6 in - let new_size := (mod_stk_len m) in - let ram := mk_ram new_size (mod_stk m) en u_en addr wr_en d_in d_out ltac:(eauto using ram_wf) in - let tc := transf_code (mod_st m) ram (mod_datapath m) (mod_controllogic m) in - match mod_ram m with - | None => - mkmodule m.(mod_params) - (fst tc) - (snd tc) - (mod_entrypoint m) - (mod_st m) - (mod_stk m) - (mod_stk_len m) - (mod_finish m) - (mod_return m) - (mod_start m) - (mod_reset m) - (mod_clk m) - (AssocMap.set u_en (None, VScalar 1) - (AssocMap.set en (None, VScalar 1) - (AssocMap.set wr_en (None, VScalar 1) - (AssocMap.set d_out (None, VScalar 32) - (AssocMap.set d_in (None, VScalar 32) - (AssocMap.set addr (None, VScalar 32) m.(mod_scldecls))))))) - (AssocMap.set m.(mod_stk) - (None, VArray 32 (2 ^ Nat.log2_up new_size))%nat m.(mod_arrdecls)) - (Some ram) - _ (mod_ordering_wf m) _ (mod_params_wf m) - | _ => m - end). - eapply transf_code_wf. apply (mod_wf m). destruct tc eqn:?; simpl. - rewrite <- Heqp. intuition. - inversion 1; subst. auto using module_ram_wf'. -Defined. - -Definition transf_fundef := transf_fundef transf_module. - -Definition transf_program (p : program) := - transform_program transf_fundef p. - -Inductive match_assocmaps : positive -> assocmap -> assocmap -> Prop := - match_assocmap: forall p rs rs', - (forall r, r < p -> rs!r = rs'!r) -> - match_assocmaps p rs rs'. - -Inductive match_arrs : assocmap_arr -> assocmap_arr -> Prop := -| match_assocmap_arr_intro: forall ar ar', - (forall s arr, - ar ! s = Some arr -> - exists arr', - ar' ! s = Some arr' - /\ (forall addr, - array_get_error addr arr = array_get_error addr arr') - /\ arr_length arr = arr_length arr')%nat -> - (forall s, ar ! s = None -> ar' ! s = None) -> - match_arrs ar ar'. - -Inductive match_arrs_size : assocmap_arr -> assocmap_arr -> Prop := - match_arrs_size_intro : - forall nasa basa, - (forall s arr, - nasa ! s = Some arr -> - exists arr', - basa ! s = Some arr' /\ arr_length arr = arr_length arr') -> - (forall s arr, - basa ! s = Some arr -> - exists arr', - nasa ! s = Some arr' /\ arr_length arr = arr_length arr') -> - (forall s, basa ! s = None <-> nasa ! s = None) -> - match_arrs_size nasa basa. - -Definition match_empty_size (m : module) (ar : assocmap_arr) : Prop := - match_arrs_size (empty_stack m) ar. -#[local] Hint Unfold match_empty_size : mgen. - -Definition disable_ram (ram: option ram) (asr : assocmap_reg) : Prop := - match ram with - | Some r => Int.eq (asr # ((ram_en r), 32)) (asr # ((ram_u_en r), 32)) = true - | None => True - end. - -Inductive match_stackframes : stackframe -> stackframe -> Prop := - match_stackframe_intro : - forall r m pc asr asa asr' asa' - (DISABLE_RAM: disable_ram (mod_ram (transf_module m)) asr') - (MATCH_ASSOC: match_assocmaps (max_reg_module m + 1) asr asr') - (MATCH_ARRS: match_arrs asa asa') - (MATCH_EMPT1: match_empty_size m asa) - (MATCH_EMPT2: match_empty_size m asa') - (MATCH_RES: r < mod_st m), - match_stackframes (Stackframe r m pc asr asa) - (Stackframe r (transf_module m) pc asr' asa'). - -Inductive match_states : state -> state -> Prop := -| match_state : - forall res res' m st asr asr' asa asa' - (ASSOC: match_assocmaps (max_reg_module m + 1) asr asr') - (ARRS: match_arrs asa asa') - (STACKS: list_forall2 match_stackframes res res') - (ARRS_SIZE: match_empty_size m asa) - (ARRS_SIZE2: match_empty_size m asa') - (DISABLE_RAM: disable_ram (mod_ram (transf_module m)) asr'), - match_states (State res m st asr asa) - (State res' (transf_module m) st asr' asa') -| match_returnstate : - forall res res' i - (STACKS: list_forall2 match_stackframes res res'), - match_states (Returnstate res i) (Returnstate res' i) -| match_initial_call : - forall m, - match_states (Callstate nil m nil) - (Callstate nil (transf_module m) nil). -#[local] Hint Constructors match_states : htlproof. - -Definition empty_stack_ram r := - AssocMap.set (ram_mem r) (Array.arr_repeat None (ram_size r)) (AssocMap.empty arr). - -Definition empty_stack' len st := - AssocMap.set st (Array.arr_repeat None len) (AssocMap.empty arr). - -Definition match_empty_size' l s (ar : assocmap_arr) : Prop := - match_arrs_size (empty_stack' l s) ar. -#[local] Hint Unfold match_empty_size : mgen. - -Definition merge_reg_assocs r := - Verilog.mkassociations (Verilog.merge_regs (assoc_nonblocking r) (assoc_blocking r)) empty_assocmap. - -Definition merge_arr_assocs st len r := - Verilog.mkassociations (Verilog.merge_arrs (assoc_nonblocking r) (assoc_blocking r)) (empty_stack' len st). - -Inductive match_reg_assocs : positive -> reg_associations -> reg_associations -> Prop := - match_reg_association: - forall p rab rab' ran ran', - match_assocmaps p rab rab' -> - match_assocmaps p ran ran' -> - match_reg_assocs p (mkassociations rab ran) (mkassociations rab' ran'). - -Inductive match_arr_assocs : arr_associations -> arr_associations -> Prop := - match_arr_association: - forall rab rab' ran ran', - match_arrs rab rab' -> - match_arrs ran ran' -> - match_arr_assocs (mkassociations rab ran) (mkassociations rab' ran'). - -Ltac mgen_crush := crush; eauto with mgen. - -Lemma match_assocmaps_equiv : forall p a, match_assocmaps p a a. -Proof. constructor; auto. Qed. -#[local] Hint Resolve match_assocmaps_equiv : mgen. - -Lemma match_arrs_equiv : forall a, match_arrs a a. -Proof. econstructor; mgen_crush. Qed. -#[local] Hint Resolve match_arrs_equiv : mgen. - -Lemma match_reg_assocs_equiv : forall p a, match_reg_assocs p a a. -Proof. destruct a; constructor; mgen_crush. Qed. -#[local] Hint Resolve match_reg_assocs_equiv : mgen. - -Lemma match_arr_assocs_equiv : forall a, match_arr_assocs a a. -Proof. destruct a; constructor; mgen_crush. Qed. -#[local] Hint Resolve match_arr_assocs_equiv : mgen. - -Lemma match_arrs_size_equiv : forall a, match_arrs_size a a. -Proof. intros; repeat econstructor; eauto. Qed. -#[local] Hint Resolve match_arrs_size_equiv : mgen. - -Lemma match_stacks_equiv : - forall m s l, - mod_stk m = s -> - mod_stk_len m = l -> - empty_stack' l s = empty_stack m. -Proof. unfold empty_stack', empty_stack; mgen_crush. Qed. -#[local] Hint Rewrite match_stacks_equiv : mgen. - -Lemma match_assocmaps_max1 : - forall p p' a b, - match_assocmaps (Pos.max p' p) a b -> - match_assocmaps p a b. -Proof. - intros. inv H. constructor. intros. - apply H0. lia. -Qed. -#[local] Hint Resolve match_assocmaps_max1 : mgen. - -Lemma match_assocmaps_max2 : - forall p p' a b, - match_assocmaps (Pos.max p p') a b -> - match_assocmaps p a b. -Proof. - intros. inv H. constructor. intros. - apply H0. lia. -Qed. -#[local] Hint Resolve match_assocmaps_max2 : mgen. - -Lemma match_assocmaps_ge : - forall p p' a b, - match_assocmaps p' a b -> - p <= p' -> - match_assocmaps p a b. -Proof. - intros. inv H. constructor. intros. - apply H1. lia. -Qed. -#[local] Hint Resolve match_assocmaps_ge : mgen. - -Lemma match_reg_assocs_max1 : - forall p p' a b, - match_reg_assocs (Pos.max p' p) a b -> - match_reg_assocs p a b. -Proof. intros; inv H; econstructor; mgen_crush. Qed. -#[local] Hint Resolve match_reg_assocs_max1 : mgen. - -Lemma match_reg_assocs_max2 : - forall p p' a b, - match_reg_assocs (Pos.max p p') a b -> - match_reg_assocs p a b. -Proof. intros; inv H; econstructor; mgen_crush. Qed. -#[local] Hint Resolve match_reg_assocs_max2 : mgen. - -Lemma match_reg_assocs_ge : - forall p p' a b, - match_reg_assocs p' a b -> - p <= p' -> - match_reg_assocs p a b. -Proof. intros; inv H; econstructor; mgen_crush. Qed. -#[local] Hint Resolve match_reg_assocs_ge : mgen. - -Definition forall_ram (P: reg -> Prop) ram := - P (ram_en ram) - /\ P (ram_u_en ram) - /\ P (ram_addr ram) - /\ P (ram_wr_en ram) - /\ P (ram_d_in ram) - /\ P (ram_d_out ram). - -Lemma forall_ram_lt : - forall m r, - (mod_ram m) = Some r -> - forall_ram (fun x => x < max_reg_module m + 1) r. -Proof. - assert (X: forall a b c, a < b + 1 -> a < Pos.max c b + 1) by lia. - unfold forall_ram; simplify; unfold HTL.max_reg_module; repeat apply X; - unfold HTL.max_reg_ram; rewrite H; try lia. -Qed. -#[local] Hint Resolve forall_ram_lt : mgen. - -Definition exec_all d_s c_s rs1 ar1 rs3 ar3 := - exists f rs2 ar2, - stmnt_runp f rs1 ar1 c_s rs2 ar2 - /\ stmnt_runp f rs2 ar2 d_s rs3 ar3. - -Definition exec_all_ram r d_s c_s rs1 ar1 rs4 ar4 := - exists f rs2 ar2 rs3 ar3, - stmnt_runp f rs1 ar1 c_s rs2 ar2 - /\ stmnt_runp f rs2 ar2 d_s rs3 ar3 - /\ exec_ram (merge_reg_assocs rs3) (merge_arr_assocs (ram_mem r) (ram_size r) ar3) (Some r) rs4 ar4. - -(* Lemma merge_reg_idempotent : *) -(* forall rs p, *) -(* match_reg_assocs p (merge_reg_assocs (merge_reg_assocs rs)) (merge_reg_assocs rs). *) -(* Proof. intros. unfold merge_reg_assocs. cbn. unfold merge_regs. *) -(* #[global] Hint Rewrite merge_reg_idempotent : mgen. *) - -Lemma merge_arr_idempotent : - forall ar st len v v', - (assoc_nonblocking ar)!st = Some v -> - (assoc_blocking ar)!st = Some v' -> - arr_length v' = len -> - arr_length v = len -> - (assoc_blocking (merge_arr_assocs st len (merge_arr_assocs st len ar)))!st - = (assoc_blocking (merge_arr_assocs st len ar))!st - /\ (assoc_nonblocking (merge_arr_assocs st len (merge_arr_assocs st len ar)))!st - = (assoc_nonblocking (merge_arr_assocs st len ar))!st. -Proof. - split; simplify; eauto. - unfold merge_arrs. - rewrite AssocMap.gcombine by reflexivity. - unfold empty_stack'. - (* rewrite AssocMap.gss. *) - (* setoid_rewrite merge_arr_empty2; auto. *) - (* rewrite AssocMap.gcombine by reflexivity. *) - (* unfold merge_arr, arr. *) - (* rewrite H. rewrite H0. auto. *) - (* unfold combine. *) - (* simplify. *) - (* rewrite list_combine_length. *) - (* rewrite (arr_wf v). rewrite (arr_wf v'). *) - (* lia. *) -(* Qed. *) Admitted. - -Lemma empty_arr : - forall m s, - (exists l, (empty_stack m) ! s = Some (arr_repeat None l)) - \/ (empty_stack m) ! s = None. -Proof. - unfold empty_stack. simplify. - destruct (Pos.eq_dec s (mod_stk m)); subst. - left. eexists. apply AssocMap.gss. - right. rewrite AssocMap.gso; auto. -Qed. - -Lemma merge_arr_empty': - forall m ar s v, - match_empty_size m ar -> - (merge_arrs (empty_stack m) ar) ! s = v -> - ar ! s = v. -Proof. - inversion 1; subst. - pose proof (empty_arr m s). - simplify. - destruct (merge_arrs (empty_stack m) ar) ! s eqn:?; subst. - - inv H3. inv H4. - learn H3 as H5. apply H0 in H5. inv H5. simplify. - unfold merge_arrs in *. (* rewrite AssocMap.gcombine in Heqo; auto. *) - (* rewrite H3 in Heqo. erewrite merge_arr_empty2 in Heqo. auto. eauto. *) - (* rewrite list_repeat_len in H6. auto. *) - (* learn H4 as H6. apply H2 in H6. *) - (* unfold merge_arrs in *. rewrite AssocMap.gcombine in Heqo; auto. *) - (* rewrite H4 in Heqo. unfold Verilog.arr in *. rewrite H6 in Heqo. *) - (* discriminate. *) - (* - inv H3. inv H4. learn H3 as H4. apply H0 in H4. inv H4. simplify. *) - (* rewrite list_repeat_len in H6. *) - (* unfold merge_arrs in *. rewrite AssocMap.gcombine in Heqo; auto. rewrite H3 in Heqo. *) - (* unfold Verilog.arr in *. rewrite H4 in Heqo. *) - (* discriminate. *) - (* apply H2 in H4; auto. *) -(* Qed. *) Admitted. - -Lemma merge_arr_empty : - forall m ar ar', - match_empty_size m ar -> - match_arrs ar ar' -> - match_arrs (merge_arrs (empty_stack m) ar) ar'. -Proof. - inversion 1; subst; inversion 1; subst; - econstructor; intros; apply merge_arr_empty' in H6; auto. -Qed. -#[local] Hint Resolve merge_arr_empty : mgen. - -Lemma merge_arr_empty'': - forall m ar s v, - match_empty_size m ar -> - ar ! s = v -> - (merge_arrs (empty_stack m) ar) ! s = v. -Proof. - inversion 1; subst. - pose proof (empty_arr m s). - simplify. - destruct (merge_arrs (empty_stack m) ar) ! s eqn:?; subst. - - inv H3. inv H4. - learn H3 as H5. apply H0 in H5. inv H5. simplify. - unfold merge_arrs in *. (* rewrite AssocMap.gcombine in Heqo; auto. *) -(* rewrite H3 in Heqo. erewrite merge_arr_empty2 in Heqo. auto. eauto. *) -(* rewrite list_repeat_len in H6. auto. *) -(* learn H4 as H6. apply H2 in H6. *) -(* unfold merge_arrs in *. rewrite AssocMap.gcombine in Heqo; auto. *) -(* rewrite H4 in Heqo. unfold Verilog.arr in *. rewrite H6 in Heqo. *) -(* discriminate. *) -(* - inv H3. inv H4. learn H3 as H4. apply H0 in H4. inv H4. simplify. *) -(* rewrite list_repeat_len in H6. *) -(* unfold merge_arrs in *. rewrite AssocMap.gcombine in Heqo; auto. rewrite H3 in Heqo. *) -(* unfold Verilog.arr in *. rewrite H4 in Heqo. *) -(* discriminate. *) -(* apply H2 in H4; auto. *) -(* Qed. *) Admitted. - -Lemma merge_arr_empty_match : - forall m ar, - match_empty_size m ar -> - match_empty_size m (merge_arrs (empty_stack m) ar). -Proof. - inversion 1; subst. - constructor; simplify. - learn H3 as H4. apply H0 in H4. inv H4. simplify. - eexists; split; eauto. apply merge_arr_empty''; eauto. - apply merge_arr_empty' in H3; auto. - split; simplify. - unfold merge_arrs in H3. (* rewrite AssocMap.gcombine in H3; auto. *) -(* unfold merge_arr in *. *) -(* destruct (empty_stack m) ! s eqn:?; *) -(* destruct ar ! s; try discriminate; eauto. *) -(* apply merge_arr_empty''; auto. apply H2 in H3; auto. *) -(* Qed. *) Admitted. -#[local] Hint Resolve merge_arr_empty_match : mgen. - -Definition ram_present {A: Type} ar r v v' := - (assoc_nonblocking ar)!(ram_mem r) = Some v - /\ @arr_length A v = ram_size r - /\ (assoc_blocking ar)!(ram_mem r) = Some v' - /\ arr_length v' = ram_size r. - -Lemma find_assoc_get : - forall rs r trs, - rs ! r = trs ! r -> - rs # r = trs # r. -Proof. - intros; unfold find_assocmap, AssocMapExt.get_default; rewrite H; auto. -Qed. -#[local] Hint Resolve find_assoc_get : mgen. - -Lemma find_assoc_get2 : - forall rs p r v trs, - (forall r, r < p -> rs ! r = trs ! r) -> - r < p -> - rs # r = v -> - trs # r = v. -Proof. - intros; unfold find_assocmap, AssocMapExt.get_default; rewrite <- H; auto. -Qed. -#[local] Hint Resolve find_assoc_get2 : mgen. - -Lemma get_assoc_gt : - forall A (rs : AssocMap.t A) p r v trs, - (forall r, r < p -> rs ! r = trs ! r) -> - r < p -> - rs ! r = v -> - trs ! r = v. -Proof. intros. rewrite <- H; auto. Qed. -#[local] Hint Resolve get_assoc_gt : mgen. - -Lemma expr_runp_matches : - forall f rs ar e v, - expr_runp f rs ar e v -> - forall trs tar, - match_assocmaps (max_reg_expr e + 1) rs trs -> - match_arrs ar tar -> - expr_runp f trs tar e v. -Proof. - induction 1. - - intros. econstructor. - - intros. econstructor. inv H0. symmetry. - apply find_assoc_get. - apply H2. crush. - - intros. econstructor. apply IHexpr_runp; eauto. - inv H1. constructor. simplify. - assert (forall a b c, a < b + 1 -> a < Pos.max c b + 1) by lia. - eapply H4 in H1. eapply H3 in H1. auto. - inv H2. - unfold arr_assocmap_lookup in *. - destruct (stack ! r) eqn:?; [|discriminate]. - inv H1. - inv H0. - eapply H3 in Heqo. inv Heqo. inv H0. - unfold arr in *. rewrite H1. inv H5. - rewrite H0. auto. - - intros. econstructor; eauto. eapply IHexpr_runp1; eauto. - econstructor. inv H2. intros. - assert (forall a b c, a < b + 1 -> a < Pos.max b c + 1) by lia. - simplify. - eapply H5 in H2. apply H4 in H2. auto. - apply IHexpr_runp2; eauto. - econstructor. inv H2. intros. - assert (forall a b c, a < b + 1 -> a < Pos.max c b + 1) by lia. - simplify. eapply H5 in H2. apply H4 in H2. auto. - - intros. econstructor; eauto. - - intros. econstructor; eauto. apply IHexpr_runp1; eauto. - constructor. inv H2. intros. simplify. - assert (forall a b c, a < b + 1 -> a < Pos.max b c + 1) by lia. - eapply H5 in H2. apply H4 in H2. auto. - apply IHexpr_runp2; eauto. simplify. - assert (forall a b c d, a < b + 1 -> a < Pos.max c (Pos.max b d) + 1) by lia. - constructor. intros. eapply H4 in H5. inv H2. apply H6 in H5. auto. - - intros. eapply erun_Vternary_false. apply IHexpr_runp1; eauto. constructor. inv H2. - intros. simplify. assert (forall a b c, a < b + 1 -> a < Pos.max b c + 1) by lia. - eapply H5 in H2. apply H4 in H2. auto. - apply IHexpr_runp2; eauto. econstructor. inv H2. simplify. - assert (forall a b c d, a < b + 1 -> a < Pos.max c (Pos.max d b) + 1) by lia. - eapply H5 in H2. apply H4 in H2. auto. auto. -Qed. -#[local] Hint Resolve expr_runp_matches : mgen. - -Lemma expr_runp_matches2 : - forall f rs ar e v p, - expr_runp f rs ar e v -> - max_reg_expr e < p -> - forall trs tar, - match_assocmaps p rs trs -> - match_arrs ar tar -> - expr_runp f trs tar e v. -Proof. - intros. eapply expr_runp_matches; eauto. - assert (max_reg_expr e + 1 <= p) by lia. - mgen_crush. -Qed. -#[local] Hint Resolve expr_runp_matches2 : mgen. - -Lemma match_assocmaps_gss : - forall p rab rab' r rhsval, - match_assocmaps p rab rab' -> - match_assocmaps p rab # r <- rhsval rab' # r <- rhsval. -Proof. - intros. inv H. econstructor. - intros. - unfold find_assocmap. unfold AssocMapExt.get_default. - destruct (Pos.eq_dec r r0); subst. - repeat rewrite PTree.gss; auto. - repeat rewrite PTree.gso; auto. -Qed. -#[local] Hint Resolve match_assocmaps_gss : mgen. - -Lemma match_assocmaps_gt : - forall p s ra ra' v, - p <= s -> - match_assocmaps p ra ra' -> - match_assocmaps p ra (ra' # s <- v). -Proof. - intros. inv H0. constructor. - intros. rewrite AssocMap.gso; try lia. apply H1; auto. -Qed. -#[local] Hint Resolve match_assocmaps_gt : mgen. - -Lemma match_reg_assocs_block : - forall p rab rab' r rhsval, - match_reg_assocs p rab rab' -> - match_reg_assocs p (block_reg r rab rhsval) (block_reg r rab' rhsval). -Proof. inversion 1; econstructor; eauto with mgen. Qed. -#[local] Hint Resolve match_reg_assocs_block : mgen. - -Lemma match_reg_assocs_nonblock : - forall p rab rab' r rhsval, - match_reg_assocs p rab rab' -> - match_reg_assocs p (nonblock_reg r rab rhsval) (nonblock_reg r rab' rhsval). -Proof. inversion 1; econstructor; eauto with mgen. Qed. -#[local] Hint Resolve match_reg_assocs_nonblock : mgen. - -Lemma some_inj : - forall A (x: A) y, - Some x = Some y -> - x = y. -Proof. inversion 1; auto. Qed. - -Lemma arrs_present : - forall r i v ar arr, - (arr_assocmap_set r i v ar) ! r = Some arr -> - exists b, ar ! r = Some b. -Proof. - intros. unfold arr_assocmap_set in *. - destruct ar!r eqn:?. - rewrite AssocMap.gss in H. - inv H. eexists. auto. rewrite Heqo in H. discriminate. -Qed. - -Ltac inv_exists := - match goal with - | H: exists _, _ |- _ => inv H - end. - -Lemma array_get_error_bound_gt : - forall A i (a : Array A), - (i >= arr_length a)%nat -> - array_get_error i a = None. -Proof. - intros. unfold array_get_error. - apply nth_error_None. destruct a. simplify. - lia. -Qed. -#[local] Hint Resolve array_get_error_bound_gt : mgen. - -Lemma array_get_error_each : - forall A addr i (v : A) a x, - arr_length a = arr_length x -> - array_get_error addr a = array_get_error addr x -> - array_get_error addr (array_set i v a) = array_get_error addr (array_set i v x). -Proof. - intros. - destruct (Nat.eq_dec addr i); subst. - destruct (lt_dec i (arr_length a)). - repeat rewrite array_get_error_set_bound; auto. - rewrite <- H. auto. - apply Nat.nlt_ge in n. - repeat rewrite array_get_error_bound_gt; auto. - rewrite <- array_set_len. rewrite <- H. lia. - repeat rewrite array_gso; auto. -Qed. -#[local] Hint Resolve array_get_error_each : mgen. - -Lemma match_arrs_gss : - forall ar ar' r v i, - match_arrs ar ar' -> - match_arrs (arr_assocmap_set r i v ar) (arr_assocmap_set r i v ar'). -Proof. - Ltac mag_tac := - match goal with - | H: ?ar ! _ = Some _, H2: forall s arr, ?ar ! s = Some arr -> _ |- _ => - let H3 := fresh "H" in - learn H as H3; apply H2 in H3; inv_exists; simplify - | H: ?ar ! _ = None, H2: forall s, ?ar ! s = None -> _ |- _ => - let H3 := fresh "H" in - learn H as H3; apply H2 in H3; inv_exists; simplify - | H: ?ar ! _ = _ |- context[match ?ar ! _ with _ => _ end] => - unfold Verilog.arr in *; rewrite H - | H: ?ar ! _ = _, H2: context[match ?ar ! _ with _ => _ end] |- _ => - unfold Verilog.arr in *; rewrite H in H2 - | H: context[(_ # ?s <- _) ! ?s] |- _ => rewrite AssocMap.gss in H - | H: context[(_ # ?r <- _) ! ?s], H2: ?r <> ?s |- _ => rewrite AssocMap.gso in H; auto - | |- context[(_ # ?s <- _) ! ?s] => rewrite AssocMap.gss - | H: ?r <> ?s |- context[(_ # ?r <- _) ! ?s] => rewrite AssocMap.gso; auto - end. - intros. - inv H. econstructor; simplify. - destruct (Pos.eq_dec r s); subst. - - unfold arr_assocmap_set, Verilog.arr in *. - destruct ar!s eqn:?. - mag_tac. - econstructor; simplify. - repeat mag_tac; auto. - intros. repeat mag_tac. simplify. - apply array_get_error_each; auto. - repeat mag_tac; crush. - repeat mag_tac; crush. - - unfold arr_assocmap_set in *. - destruct ar ! r eqn:?. rewrite AssocMap.gso in H; auto. - apply H0 in Heqo. apply H0 in H. repeat inv_exists. simplify. - econstructor. unfold Verilog.arr in *. rewrite H3. simplify. - rewrite AssocMap.gso; auto. eauto. intros. auto. auto. - apply H1 in Heqo. apply H0 in H. repeat inv_exists; simplify. - econstructor. unfold Verilog.arr in *. rewrite Heqo. simplify; eauto. - - destruct (Pos.eq_dec r s); unfold arr_assocmap_set, Verilog.arr in *; simplify; subst. - destruct ar!s eqn:?; repeat mag_tac; crush. - apply H1 in H. mag_tac; crush. - destruct ar!r eqn:?; repeat mag_tac; crush. - apply H1 in Heqo. repeat mag_tac; auto. -Qed. -#[local] Hint Resolve match_arrs_gss : mgen. - -Lemma match_arr_assocs_block : - forall rab rab' r i rhsval, - match_arr_assocs rab rab' -> - match_arr_assocs (block_arr r i rab rhsval) (block_arr r i rab' rhsval). -Proof. inversion 1; econstructor; eauto with mgen. Qed. -#[local] Hint Resolve match_arr_assocs_block : mgen. - -Lemma match_arr_assocs_nonblock : - forall rab rab' r i rhsval, - match_arr_assocs rab rab' -> - match_arr_assocs (nonblock_arr r i rab rhsval) (nonblock_arr r i rab' rhsval). -Proof. inversion 1; econstructor; eauto with mgen. Qed. -#[local] Hint Resolve match_arr_assocs_nonblock : mgen. - -Lemma match_states_same : - forall f rs1 ar1 c rs2 ar2 p, - stmnt_runp f rs1 ar1 c rs2 ar2 -> - max_reg_stmnt c < p -> - forall trs1 tar1, - match_reg_assocs p rs1 trs1 -> - match_arr_assocs ar1 tar1 -> - exists trs2 tar2, - stmnt_runp f trs1 tar1 c trs2 tar2 - /\ match_reg_assocs p rs2 trs2 - /\ match_arr_assocs ar2 tar2. -Proof. - Ltac match_states_same_facts := - match goal with - | H: match_reg_assocs _ _ _ |- _ => - let H2 := fresh "H" in - learn H as H2; inv H2 - | H: match_arr_assocs _ _ |- _ => - let H2 := fresh "H" in - learn H as H2; inv H2 - | H1: context[exists _, _], H2: context[exists _, _] |- _ => - learn H1; learn H2; - exploit H1; mgen_crush; exploit H2; mgen_crush - | H1: context[exists _, _] |- _ => - learn H1; exploit H1; mgen_crush - end. - Ltac match_states_same_tac := - match goal with - | |- exists _, _ => econstructor - | |- _ < _ => lia - | H: context[_ <> _] |- stmnt_runp _ _ _ (Vcase _ (Stmntcons _ _ _) _) _ _ => - eapply stmnt_runp_Vcase_nomatch - | |- stmnt_runp _ _ _ (Vcase _ (Stmntcons _ _ _) _) _ _ => - eapply stmnt_runp_Vcase_match - | H: valueToBool _ = false |- stmnt_runp _ _ _ _ _ _ => - eapply stmnt_runp_Vcond_false - | |- stmnt_runp _ _ _ _ _ _ => econstructor - | |- expr_runp _ _ _ _ _ => eapply expr_runp_matches2 - end; mgen_crush; try lia. - induction 1; simplify; repeat match_states_same_facts; - try destruct_match; try solve [repeat match_states_same_tac]. - - inv H. exists (block_reg r {| assoc_blocking := rab'; assoc_nonblocking := ran' |} rhsval); - repeat match_states_same_tac; econstructor. - - exists (nonblock_reg r {| assoc_blocking := rab'; assoc_nonblocking := ran' |} rhsval); - repeat match_states_same_tac; inv H; econstructor. - - econstructor. exists (block_arr r i {| assoc_blocking := rab'0; assoc_nonblocking := ran'0 |} rhsval). - simplify; repeat match_states_same_tac. inv H. econstructor. - repeat match_states_same_tac. eauto. mgen_crush. - - econstructor. exists (nonblock_arr r i {| assoc_blocking := rab'0; assoc_nonblocking := ran'0 |} rhsval). - simplify; repeat match_states_same_tac. inv H. econstructor. - repeat match_states_same_tac. eauto. mgen_crush. -Qed. - -Lemma match_reg_assocs_merge : - forall p rs rs', - match_reg_assocs p rs rs' -> - match_reg_assocs p (merge_reg_assocs rs) (merge_reg_assocs rs'). -Proof. - inversion 1. - econstructor; econstructor; crush. inv H3. inv H. - inv H7. inv H9. - unfold merge_regs. - destruct (ran!r) eqn:?; destruct (rab!r) eqn:?. - erewrite AssocMapExt.merge_correct_1; eauto. - erewrite AssocMapExt.merge_correct_1; eauto. - rewrite <- H2; eauto. - erewrite AssocMapExt.merge_correct_1; eauto. - erewrite AssocMapExt.merge_correct_1; eauto. - rewrite <- H2; eauto. - erewrite AssocMapExt.merge_correct_2; eauto. - erewrite AssocMapExt.merge_correct_2; eauto. - rewrite <- H2; eauto. - rewrite <- H; eauto. - erewrite AssocMapExt.merge_correct_3; eauto. - erewrite AssocMapExt.merge_correct_3; eauto. - rewrite <- H2; eauto. - rewrite <- H; eauto. -Qed. -#[local] Hint Resolve match_reg_assocs_merge : mgen. - -Lemma transf_not_changed : - forall state ram n d c i d_s c_s, - (forall e1 e2 r, d_s <> Vnonblock (Vvari r e1) e2) -> - (forall e1 e2 r, d_s <> Vnonblock e1 (Vvari r e2)) -> - d!i = Some d_s -> - c!i = Some c_s -> - transf_maps state ram (i, n) (d, c) = (d, c). -Proof. intros; unfold transf_maps; repeat destruct_match; mgen_crush. Qed. - -Lemma transf_not_changed_neq : - forall state ram n d c d' c' i a d_s c_s, - transf_maps state ram (a, n) (d, c) = (d', c') -> - d!i = Some d_s -> - c!i = Some c_s -> - a <> i -> n <> i -> - d'!i = Some d_s /\ c'!i = Some c_s. -Proof. - unfold transf_maps; intros; repeat destruct_match; mgen_crush; - match goal with [ H: (_, _) = (_, _) |- _ ] => inv H end; - repeat (rewrite AssocMap.gso; auto). -Qed. - -Lemma forall_gt : - forall l, Forall (Pos.ge (fold_right Pos.max 1 l)) l. -Proof. - induction l; auto. - constructor. inv IHl; simplify; lia. - simplify. destruct (Pos.max_dec a (fold_right Pos.max 1 l)). - rewrite e. apply Pos.max_l_iff in e. apply Pos.le_ge in e. - apply Forall_forall. rewrite Forall_forall in IHl. - intros. - assert (X: forall a b c, a >= c -> c >= b -> a >= b) by lia. - apply X with (b := x) in e; auto. - rewrite e; auto. -Qed. - -Lemma max_index_list : - forall A (l : list (positive * A)) i d_s, - In (i, d_s) l -> - list_norepet (map fst l) -> - i <= fold_right Pos.max 1 (map fst l). -Proof. - induction l; crush. - inv H. inv H0. simplify. lia. - inv H0. - let H := fresh "H" in - assert (H: forall a b c, c <= b -> c <= Pos.max a b) by lia; - apply H; eapply IHl; eauto. -Qed. - -Lemma max_index : - forall A d i (d_s: A), d ! i = Some d_s -> i <= max_pc d. -Proof. - unfold max_pc; - eauto using max_index_list, - PTree.elements_correct, PTree.elements_keys_norepet. -Qed. - -Lemma max_index_2' : - forall l i, i > fold_right Pos.max 1 l -> Forall (Pos.gt i) l. -Proof. induction l; crush; constructor; [|apply IHl]; lia. Qed. - -Lemma max_index_2'' : - forall l i, Forall (Pos.gt i) l -> ~ In i l. -Proof. - induction l; crush. unfold not in *. - intros. inv H0. inv H. lia. eapply IHl. - inv H. apply H4. auto. -Qed. - -Lemma elements_correct_none : - forall A am k, - ~ List.In k (List.map (@fst _ A) (AssocMap.elements am)) -> - AssocMap.get k am = None. -Proof. - intros. apply AssocMapExt.elements_correct' in H. unfold not in *. - destruct am ! k eqn:?; auto. exfalso. apply H. eexists. auto. -Qed. -#[local] Hint Resolve elements_correct_none : assocmap. - -Lemma max_index_2 : - forall A (d: AssocMap.t A) i, i > max_pc d -> d ! i = None. -Proof. - intros. unfold max_pc in *. apply max_index_2' in H. - apply max_index_2'' in H. apply elements_correct_none. auto. -Qed. - -Definition match_prog (p: program) (tp: program) := - Linking.match_program (fun cu f tf => tf = transf_fundef f) eq p tp. - -Lemma transf_program_match: - forall p, match_prog p (transf_program p). -Proof. - intros. unfold transf_program, match_prog. - apply Linking.match_transform_program. -Qed. - -Lemma exec_all_Vskip : - forall rs ar, - exec_all Vskip Vskip rs ar rs ar. -Proof. - unfold exec_all. - intros. repeat econstructor. - Unshelve. unfold fext. exact tt. -Qed. - -Lemma match_assocmaps_trans: - forall p rs1 rs2 rs3, - match_assocmaps p rs1 rs2 -> - match_assocmaps p rs2 rs3 -> - match_assocmaps p rs1 rs3. -Proof. - intros. inv H. inv H0. econstructor; eauto. - intros. rewrite H1 by auto. auto. -Qed. - -Lemma match_reg_assocs_trans: - forall p rs1 rs2 rs3, - match_reg_assocs p rs1 rs2 -> - match_reg_assocs p rs2 rs3 -> - match_reg_assocs p rs1 rs3. -Proof. - intros. inv H. inv H0. - econstructor; eapply match_assocmaps_trans; eauto. -Qed. - -Lemma empty_arrs_match : - forall m, - match_arrs (empty_stack m) (empty_stack (transf_module m)). -Proof. - intros; - unfold empty_stack, transf_module; repeat destruct_match; mgen_crush. -Qed. -#[local] Hint Resolve empty_arrs_match : mgen. - -Lemma max_module_stmnts : - forall m, - Pos.max (max_stmnt_tree (mod_controllogic m)) - (max_stmnt_tree (mod_datapath m)) < max_reg_module m + 1. -Proof. intros. unfold max_reg_module. lia. Qed. -#[local] Hint Resolve max_module_stmnts : mgen. - -Lemma transf_module_code : - forall m, - mod_ram m = None -> - transf_code (mod_st m) - {| ram_size := mod_stk_len m; - ram_mem := mod_stk m; - ram_en := max_reg_module m + 2; - ram_addr := max_reg_module m + 1; - ram_wr_en := max_reg_module m + 5; - ram_d_in := max_reg_module m + 3; - ram_d_out := max_reg_module m + 4; - ram_u_en := max_reg_module m + 6; - ram_ordering := ram_wf (max_reg_module m) |} - (mod_datapath m) (mod_controllogic m) - = ((mod_datapath (transf_module m)), mod_controllogic (transf_module m)). -Proof. unfold transf_module; intros; repeat destruct_match; crush. - apply surjective_pairing. Qed. -#[local] Hint Resolve transf_module_code : mgen. - -Lemma transf_module_code_ram : - forall m r, mod_ram m = Some r -> transf_module m = m. -Proof. unfold transf_module; intros; repeat destruct_match; crush. Qed. -#[local] Hint Resolve transf_module_code_ram : mgen. - -Lemma mod_reset_lt : forall m, mod_reset m < max_reg_module m + 1. -Proof. intros; unfold max_reg_module; lia. Qed. -#[local] Hint Resolve mod_reset_lt : mgen. - -Lemma mod_finish_lt : forall m, mod_finish m < max_reg_module m + 1. -Proof. intros; unfold max_reg_module; lia. Qed. -#[local] Hint Resolve mod_finish_lt : mgen. - -Lemma mod_return_lt : forall m, mod_return m < max_reg_module m + 1. -Proof. intros; unfold max_reg_module; lia. Qed. -#[local] Hint Resolve mod_return_lt : mgen. - -Lemma mod_start_lt : forall m, mod_start m < max_reg_module m + 1. -Proof. intros; unfold max_reg_module; lia. Qed. -#[local] Hint Resolve mod_start_lt : mgen. - -Lemma mod_stk_lt : forall m, mod_stk m < max_reg_module m + 1. -Proof. intros; unfold max_reg_module; lia. Qed. -#[local] Hint Resolve mod_stk_lt : mgen. - -Lemma mod_st_lt : forall m, mod_st m < max_reg_module m + 1. -Proof. intros; unfold max_reg_module; lia. Qed. -#[local] Hint Resolve mod_st_lt : mgen. - -Lemma mod_reset_modify : - forall m ar ar' v, - match_assocmaps (max_reg_module m + 1) ar ar' -> - ar ! (mod_reset m) = Some v -> - ar' ! (mod_reset (transf_module m)) = Some v. -Proof. - inversion 1. intros. - unfold transf_module; repeat destruct_match; simplify; - rewrite <- H0; eauto with mgen. -Qed. -#[local] Hint Resolve mod_reset_modify : mgen. - -Lemma mod_finish_modify : - forall m ar ar' v, - match_assocmaps (max_reg_module m + 1) ar ar' -> - ar ! (mod_finish m) = Some v -> - ar' ! (mod_finish (transf_module m)) = Some v. -Proof. - inversion 1. intros. - unfold transf_module; repeat destruct_match; simplify; - rewrite <- H0; eauto with mgen. -Qed. -#[local] Hint Resolve mod_finish_modify : mgen. - -Lemma mod_return_modify : - forall m ar ar' v, - match_assocmaps (max_reg_module m + 1) ar ar' -> - ar ! (mod_return m) = Some v -> - ar' ! (mod_return (transf_module m)) = Some v. -Proof. - inversion 1. intros. - unfold transf_module; repeat destruct_match; simplify; - rewrite <- H0; eauto with mgen. -Qed. -#[local] Hint Resolve mod_return_modify : mgen. - -Lemma mod_start_modify : - forall m ar ar' v, - match_assocmaps (max_reg_module m + 1) ar ar' -> - ar ! (mod_start m) = Some v -> - ar' ! (mod_start (transf_module m)) = Some v. -Proof. - inversion 1. intros. - unfold transf_module; repeat destruct_match; simplify; - rewrite <- H0; eauto with mgen. -Qed. -#[local] Hint Resolve mod_start_modify : mgen. - -Lemma mod_st_modify : - forall m ar ar' v, - match_assocmaps (max_reg_module m + 1) ar ar' -> - ar ! (mod_st m) = Some v -> - ar' ! (mod_st (transf_module m)) = Some v. -Proof. - inversion 1. intros. - unfold transf_module; repeat destruct_match; simplify; - rewrite <- H0; eauto with mgen. -Qed. -#[local] Hint Resolve mod_st_modify : mgen. - -Lemma match_arrs_read : - forall ra ra' addr v mem, - arr_assocmap_lookup ra mem addr = Some v -> - match_arrs ra ra' -> - arr_assocmap_lookup ra' mem addr = Some v. -Proof. - unfold arr_assocmap_lookup. intros. destruct_match; destruct_match; try discriminate. - inv H0. eapply H1 in Heqo0. inv Heqo0. simplify. unfold arr in *. - rewrite H in Heqo. inv Heqo. - rewrite H0. auto. - inv H0. eapply H1 in Heqo0. inv Heqo0. inv H0. unfold arr in *. - rewrite H3 in Heqo. discriminate. -Qed. -#[local] Hint Resolve match_arrs_read : mgen. - -Lemma match_reg_implies_equal : - forall ra ra' p a b c, - Int.eq (ra # a) (ra # b) = c -> - a < p -> b < p -> - match_assocmaps p ra ra' -> - Int.eq (ra' # a) (ra' # b) = c. -Proof. - unfold find_assocmap, AssocMapExt.get_default; intros. - inv H2. destruct (ra ! a) eqn:?; destruct (ra ! b) eqn:?; - repeat rewrite <- H3 by lia; rewrite Heqo; rewrite Heqo0; auto. -Qed. -#[local] Hint Resolve match_reg_implies_equal : mgen. - -Lemma exec_ram_same : - forall rs1 ar1 ram rs2 ar2 p, - exec_ram rs1 ar1 (Some ram) rs2 ar2 -> - forall_ram (fun x => x < p) ram -> - forall trs1 tar1, - match_reg_assocs p rs1 trs1 -> - match_arr_assocs ar1 tar1 -> - exists trs2 tar2, - exec_ram trs1 tar1 (Some ram) trs2 tar2 - /\ match_reg_assocs p rs2 trs2 - /\ match_arr_assocs ar2 tar2. -Proof. - Ltac exec_ram_same_facts := - match goal with - | H: match_reg_assocs _ _ _ |- _ => let H2 := fresh "H" in learn H as H2; inv H2 - | H: match_assocmaps _ _ _ |- _ => let H2 := fresh "H" in learn H as H2; inv H2 - | H: match_arr_assocs _ _ |- _ => let H2 := fresh "H" in learn H as H2; inv H2 - | H: match_arrs _ _ |- _ => let H2 := fresh "H" in learn H as H2; inv H2 - end. - inversion 1; subst; destruct ram; unfold forall_ram; simplify; repeat exec_ram_same_facts. - - repeat (econstructor; mgen_crush). - - do 2 econstructor; simplify; - [eapply exec_ram_Some_write; [ apply H1 | apply H2 | | | | | ] | | ]; - mgen_crush. - - do 2 econstructor; simplify; [eapply exec_ram_Some_read | | ]; - repeat (try econstructor; mgen_crush). -Qed. - -Lemma match_assocmaps_merge : - forall p nasr basr nasr' basr', - match_assocmaps p nasr nasr' -> - match_assocmaps p basr basr' -> - match_assocmaps p (merge_regs nasr basr) (merge_regs nasr' basr'). -Proof. - unfold merge_regs. - intros. inv H. inv H0. econstructor. - intros. - destruct nasr ! r eqn:?; destruct basr ! r eqn:?. - erewrite AssocMapExt.merge_correct_1; mgen_crush. - erewrite AssocMapExt.merge_correct_1; mgen_crush. - erewrite AssocMapExt.merge_correct_1; mgen_crush. - erewrite AssocMapExt.merge_correct_1; mgen_crush. - erewrite AssocMapExt.merge_correct_2; mgen_crush. - erewrite AssocMapExt.merge_correct_2; mgen_crush. - erewrite AssocMapExt.merge_correct_3; mgen_crush. - erewrite AssocMapExt.merge_correct_3; mgen_crush. -Qed. -#[local] Hint Resolve match_assocmaps_merge : mgen. - -Lemma list_combine_nth_error1 : - forall l l' addr v, - length l = length l' -> - nth_error l addr = Some (Some v) -> - nth_error (list_combine merge_cell l l') addr = Some (Some v). -Proof. induction l; destruct l'; destruct addr; crush. Qed. - -Lemma list_combine_nth_error2 : - forall l' l addr v, - length l = length l' -> - nth_error l addr = Some None -> - nth_error l' addr = Some (Some v) -> - nth_error (list_combine merge_cell l l') addr = Some (Some v). -Proof. induction l'; try rewrite nth_error_nil in *; destruct l; destruct addr; crush. Qed. - -Lemma list_combine_nth_error3 : - forall l l' addr, - length l = length l' -> - nth_error l addr = Some None -> - nth_error l' addr = Some None -> - nth_error (list_combine merge_cell l l') addr = Some None. -Proof. induction l; destruct l'; destruct addr; crush. Qed. - -Lemma list_combine_nth_error4 : - forall l l' addr, - length l = length l' -> - nth_error l addr = None -> - nth_error (list_combine merge_cell l l') addr = None. -Proof. induction l; destruct l'; destruct addr; crush. Qed. - -Lemma list_combine_nth_error5 : - forall l l' addr, - length l = length l' -> - nth_error l' addr = None -> - nth_error (list_combine merge_cell l l') addr = None. -Proof. induction l; destruct l'; destruct addr; crush. Qed. - -Lemma array_get_error_merge1 : - forall a a0 addr v, - arr_length a = arr_length a0 -> - array_get_error addr a = Some (Some v) -> - array_get_error addr (combine merge_cell a a0) = Some (Some v). -Proof. - unfold array_get_error, combine in *; intros; - apply list_combine_nth_error1; destruct a; destruct a0; crush. -Qed. - -Lemma array_get_error_merge2 : - forall a a0 addr v, - arr_length a = arr_length a0 -> - array_get_error addr a0 = Some (Some v) -> - array_get_error addr a = Some None -> - array_get_error addr (combine merge_cell a a0) = Some (Some v). -Proof. - unfold array_get_error, combine in *; intros; - apply list_combine_nth_error2; destruct a; destruct a0; crush. -Qed. - -Lemma array_get_error_merge3 : - forall a a0 addr, - arr_length a = arr_length a0 -> - array_get_error addr a0 = Some None -> - array_get_error addr a = Some None -> - array_get_error addr (combine merge_cell a a0) = Some None. -Proof. - unfold array_get_error, combine in *; intros; - apply list_combine_nth_error3; destruct a; destruct a0; crush. -Qed. - -Lemma array_get_error_merge4 : - forall a a0 addr, - arr_length a = arr_length a0 -> - array_get_error addr a = None -> - array_get_error addr (combine merge_cell a a0) = None. -Proof. - unfold array_get_error, combine in *; intros; - apply list_combine_nth_error4; destruct a; destruct a0; crush. -Qed. - -Lemma array_get_error_merge5 : - forall a a0 addr, - arr_length a = arr_length a0 -> - array_get_error addr a0 = None -> - array_get_error addr (combine merge_cell a a0) = None. -Proof. - unfold array_get_error, combine in *; intros; - apply list_combine_nth_error5; destruct a; destruct a0; crush. -Qed. - -Lemma match_arrs_merge' : - forall addr nasa basa arr s x x0 a a0 nasa' basa', - (AssocMap.combine merge_arr nasa basa) ! s = Some arr -> - nasa ! s = Some a -> - basa ! s = Some a0 -> - nasa' ! s = Some x0 -> - basa' ! s = Some x -> - arr_length x = arr_length x0 -> - array_get_error addr a0 = array_get_error addr x -> - arr_length a0 = arr_length x -> - array_get_error addr a = array_get_error addr x0 -> - arr_length a = arr_length x0 -> - array_get_error addr arr = array_get_error addr (combine merge_cell x0 x). -Proof. - intros. rewrite AssocMap.gcombine in H by auto. - unfold merge_arr in H. - rewrite H0 in H. rewrite H1 in H. inv H. - destruct (array_get_error addr a0) eqn:?; destruct (array_get_error addr a) eqn:?. - destruct o; destruct o0. - erewrite array_get_error_merge1; eauto. erewrite array_get_error_merge1; eauto. - rewrite <- H6 in H4. rewrite <- H8 in H4. auto. - erewrite array_get_error_merge2; eauto. erewrite array_get_error_merge2; eauto. - rewrite <- H6 in H4. rewrite <- H8 in H4. auto. - erewrite array_get_error_merge1; eauto. erewrite array_get_error_merge1; eauto. - rewrite <- H6 in H4. rewrite <- H8 in H4. auto. - erewrite array_get_error_merge3; eauto. erewrite array_get_error_merge3; eauto. - rewrite <- H6 in H4. rewrite <- H8 in H4. auto. - erewrite array_get_error_merge4; eauto. erewrite array_get_error_merge4; eauto. - rewrite <- H6 in H4. rewrite <- H8 in H4. auto. - erewrite array_get_error_merge5; eauto. erewrite array_get_error_merge5; eauto. - rewrite <- H6 in H4. rewrite <- H8 in H4. auto. - erewrite array_get_error_merge5; eauto. erewrite array_get_error_merge5; eauto. - rewrite <- H6 in H4. rewrite <- H8 in H4. auto. -Qed. - -Lemma match_arrs_merge : - forall nasa nasa' basa basa', - match_arrs_size nasa basa -> - match_arrs nasa nasa' -> - match_arrs basa basa' -> - match_arrs (merge_arrs nasa basa) (merge_arrs nasa' basa'). -Proof. - unfold merge_arrs. - intros. inv H. inv H0. inv H1. econstructor. - - intros. destruct nasa ! s eqn:?; destruct basa ! s eqn:?; unfold Verilog.arr in *. - + pose proof Heqo. apply H in Heqo. pose proof Heqo0. apply H0 in Heqo0. - repeat inv_exists. simplify. - eexists. simplify. rewrite AssocMap.gcombine; eauto. - unfold merge_arr. unfold Verilog.arr in *. rewrite H11. rewrite H12. auto. - intros. eapply match_arrs_merge'; eauto. eapply H2 in H7; eauto. - inv_exists. simplify. congruence. - rewrite AssocMap.gcombine in H1; auto. unfold merge_arr in H1. - rewrite H7 in H1. rewrite H8 in H1. inv H1. - repeat rewrite combine_length; auto. - eapply H2 in H7; eauto. inv_exists; simplify; congruence. - eapply H2 in H7; eauto. inv_exists; simplify; congruence. - + apply H2 in Heqo; inv_exists; crush. - + apply H3 in Heqo0; inv_exists; crush. - + rewrite AssocMap.gcombine in H1 by auto. unfold merge_arr in *. - rewrite Heqo in H1. rewrite Heqo0 in H1. discriminate. - - intros. rewrite AssocMap.gcombine in H1 by auto. unfold merge_arr in H1. - repeat destruct_match; crush. - rewrite AssocMap.gcombine by auto; unfold merge_arr. - apply H5 in Heqo. apply H6 in Heqo0. - unfold Verilog.arr in *. - rewrite Heqo. rewrite Heqo0. auto. -Qed. -#[local] Hint Resolve match_arrs_merge : mgen. - -Lemma match_empty_size_merge : - forall nasa2 basa2 m, - match_empty_size m nasa2 -> - match_empty_size m basa2 -> - match_empty_size m (merge_arrs nasa2 basa2). -Proof. - intros. inv H. inv H0. constructor. - simplify. unfold merge_arrs. rewrite AssocMap.gcombine by auto. - pose proof H0 as H6. apply H1 in H6. inv_exists; simplify. - pose proof H0 as H9. apply H in H9. inv_exists; simplify. - eexists. simplify. unfold merge_arr. unfold Verilog.arr in *. rewrite H6. - rewrite H9. auto. rewrite H8. symmetry. apply combine_length. congruence. - intros. - destruct (nasa2 ! s) eqn:?; destruct (basa2 ! s) eqn:?. - unfold merge_arrs in H0. rewrite AssocMap.gcombine in H0 by auto. - unfold merge_arr in *. setoid_rewrite Heqo in H0. setoid_rewrite Heqo0 in H0. inv H0. - apply H2 in Heqo. apply H4 in Heqo0. repeat inv_exists; simplify. - eexists. simplify. eauto. rewrite list_combine_length. - rewrite (arr_wf a). rewrite (arr_wf a0). lia. - unfold merge_arrs in H0. rewrite AssocMap.gcombine in H0 by auto. - unfold merge_arr in *. setoid_rewrite Heqo in H0. setoid_rewrite Heqo0 in H0. - apply H2 in Heqo. inv_exists; simplify. - econstructor; eauto. - unfold merge_arrs in H0. rewrite AssocMap.gcombine in H0 by auto. - unfold merge_arr in *. setoid_rewrite Heqo in H0. setoid_rewrite Heqo0 in H0. - inv H0. apply H4 in Heqo0. inv_exists; simplify. econstructor; eauto. - unfold merge_arrs in H0. rewrite AssocMap.gcombine in H0 by auto. - unfold merge_arr in *. setoid_rewrite Heqo in H0. setoid_rewrite Heqo0 in H0. - discriminate. - split; intros. - unfold merge_arrs in H0. rewrite AssocMap.gcombine in H0 by auto. - unfold merge_arr in *. repeat destruct_match; crush. apply H5 in Heqo0; auto. - pose proof H0. - apply H5 in H0. - apply H3 in H6. unfold merge_arrs. rewrite AssocMap.gcombine by auto. - setoid_rewrite H0. setoid_rewrite H6. auto. -Qed. -#[local] Hint Resolve match_empty_size_merge : mgen. - -Lemma match_empty_size_match : - forall m nasa2 basa2, - match_empty_size m nasa2 -> - match_empty_size m basa2 -> - match_arrs_size nasa2 basa2. -Proof. - Ltac match_empty_size_match_solve := - match goal with - | H: context[forall s arr, ?ar ! s = Some arr -> _], H2: ?ar ! _ = Some _ |- _ => - let H3 := fresh "H" in - learn H; pose proof H2 as H3; apply H in H3; repeat inv_exists - | H: context[forall s, ?ar ! s = None <-> _], H2: ?ar ! _ = None |- _ => - let H3 := fresh "H" in - learn H; pose proof H2 as H3; apply H in H3 - | H: context[forall s, _ <-> ?ar ! s = None], H2: ?ar ! _ = None |- _ => - let H3 := fresh "H" in - learn H; pose proof H2 as H3; apply H in H3 - | |- exists _, _ => econstructor - | |- _ ! _ = Some _ => eassumption - | |- _ = _ => congruence - | |- _ <-> _ => split - end; simplify. - inversion 1; inversion 1; constructor; simplify; repeat match_empty_size_match_solve. -Qed. -#[local] Hint Resolve match_empty_size_match : mgen. - -Lemma match_get_merge : - forall p ran ran' rab rab' s v, - s < p -> - match_assocmaps p ran ran' -> - match_assocmaps p rab rab' -> - (merge_regs ran rab) ! s = Some v -> - (merge_regs ran' rab') ! s = Some v. -Proof. - intros. - assert (X: match_assocmaps p (merge_regs ran rab) (merge_regs ran' rab')) by auto with mgen. - inv X. rewrite <- H3; auto. -Qed. -#[local] Hint Resolve match_get_merge : mgen. - -Ltac masrp_tac := - match goal with - | H: context[forall s arr, ?ar ! s = Some arr -> _], H2: ?ar ! _ = Some _ |- _ => - let H3 := fresh "H" in - learn H; pose proof H2 as H3; apply H in H3; repeat inv_exists - | H: context[forall s, ?ar ! s = None <-> _], H2: ?ar ! _ = None |- _ => - let H3 := fresh "H" in - learn H; pose proof H2 as H3; apply H in H3 - | H: context[forall s, _ <-> ?ar ! s = None], H2: ?ar ! _ = None |- _ => - let H3 := fresh "H" in - learn H; pose proof H2 as H3; apply H in H3 - | ra: arr_associations |- _ => - let ra1 := fresh "ran" in let ra2 := fresh "rab" in destruct ra as [ra1 ra2] - | |- _ ! _ = _ => solve [mgen_crush] - | |- _ = _ => congruence - | |- _ <> _ => lia - | H: ?ar ! ?s = _ |- context[match ?ar ! ?r with _ => _ end] => learn H; destruct (Pos.eq_dec s r); subst - | H: ?ar ! ?s = _ |- context[match ?ar ! ?s with _ => _ end] => setoid_rewrite H - | |- context[match ?ar ! ?s with _ => _ end] => destruct (ar ! s) eqn:? - | H: ?s <> ?r |- context[(_ # ?r <- _) ! ?s] => rewrite AssocMap.gso - | H: ?r <> ?s |- context[(_ # ?r <- _) ! ?s] => rewrite AssocMap.gso - | |- context[(_ # ?s <- _) ! ?s] => rewrite AssocMap.gss - | H: context[match ?ar ! ?r with _ => _ end ! ?s] |- _ => - destruct (ar ! r) eqn:?; destruct (Pos.eq_dec r s); subst - | H: ?ar ! ?s = _, H2: context[match ?ar ! ?s with _ => _ end] |- _ => - setoid_rewrite H in H2 - | H: context[match ?ar ! ?s with _ => _ end] |- _ => destruct (ar ! s) eqn:? - | H: ?s <> ?r, H2: context[(_ # ?r <- _) ! ?s] |- _ => rewrite AssocMap.gso in H2 - | H: ?r <> ?s, H2: context[(_ # ?r <- _) ! ?s] |- _ => rewrite AssocMap.gso in H2 - | H: context[(_ # ?s <- _) ! ?s] |- _ => rewrite AssocMap.gss in H - | |- context[match_empty_size] => constructor - | |- context[arr_assocmap_set] => unfold arr_assocmap_set - | H: context[arr_assocmap_set] |- _ => unfold arr_assocmap_set in H - | |- exists _, _ => econstructor - | |- _ <-> _ => split - end; simplify. - -Lemma match_empty_assocmap_set : - forall m r i rhsval asa, - match_empty_size m asa -> - match_empty_size m (arr_assocmap_set r i rhsval asa). -Proof. - inversion 1; subst; simplify. - constructor. intros. - repeat masrp_tac. - intros. do 5 masrp_tac; try solve [repeat masrp_tac]. - apply H1 in H3. inv_exists. simplify. - econstructor. simplify. apply H3. congruence. - repeat masrp_tac. destruct (Pos.eq_dec r s); subst. - rewrite AssocMap.gss in H8. discriminate. - rewrite AssocMap.gso in H8; auto. apply H2 in H8. auto. - destruct (Pos.eq_dec r s); subst. apply H1 in H5. inv_exists. simplify. - rewrite H5 in H8. discriminate. - rewrite AssocMap.gso; auto. - apply H2 in H5. auto. apply H2 in H5. auto. - Unshelve. auto. -Qed. -#[local] Hint Resolve match_empty_assocmap_set : mgen. - -Lemma match_arrs_size_stmnt_preserved : - forall m f rs1 ar1 ar2 c rs2, - stmnt_runp f rs1 ar1 c rs2 ar2 -> - match_empty_size m (assoc_nonblocking ar1) -> - match_empty_size m (assoc_blocking ar1) -> - match_empty_size m (assoc_nonblocking ar2) /\ match_empty_size m (assoc_blocking ar2). -Proof. - induction 1; inversion 1; inversion 1; eauto; simplify; try solve [repeat masrp_tac]. - subst. apply IHstmnt_runp2; apply IHstmnt_runp1; auto. - apply IHstmnt_runp2; apply IHstmnt_runp1; auto. - apply match_empty_assocmap_set. auto. - apply match_empty_assocmap_set. auto. -Qed. - -Lemma match_arrs_size_stmnt_preserved2 : - forall m f rs1 na ba na' ba' c rs2, - stmnt_runp f rs1 {| assoc_nonblocking := na; assoc_blocking := ba |} c rs2 - {| assoc_nonblocking := na'; assoc_blocking := ba' |} -> - match_empty_size m na -> - match_empty_size m ba -> - match_empty_size m na' /\ match_empty_size m ba'. -Proof. - intros. - remember ({| assoc_blocking := ba; assoc_nonblocking := na |}) as ar1. - remember ({| assoc_blocking := ba'; assoc_nonblocking := na' |}) as ar2. - assert (X1: na' = (assoc_nonblocking ar2)) by (rewrite Heqar2; auto). rewrite X1. - assert (X2: ba' = (assoc_blocking ar2)) by (rewrite Heqar2; auto). rewrite X2. - assert (X3: na = (assoc_nonblocking ar1)) by (rewrite Heqar1; auto). rewrite X3 in *. - assert (X4: ba = (assoc_blocking ar1)) by (rewrite Heqar1; auto). rewrite X4 in *. - eapply match_arrs_size_stmnt_preserved; mgen_crush. -Qed. -#[local] Hint Resolve match_arrs_size_stmnt_preserved2 : mgen. - -Lemma match_arrs_size_ram_preserved : - forall m rs1 ar1 ar2 ram rs2, - exec_ram rs1 ar1 ram rs2 ar2 -> - match_empty_size m (assoc_nonblocking ar1) -> - match_empty_size m (assoc_blocking ar1) -> - match_empty_size m (assoc_nonblocking ar2) - /\ match_empty_size m (assoc_blocking ar2). -Proof. - induction 1; inversion 1; inversion 1; subst; simplify; try solve [repeat masrp_tac]. - masrp_tac. masrp_tac. solve [repeat masrp_tac]. - masrp_tac. masrp_tac. masrp_tac. masrp_tac. masrp_tac. masrp_tac. masrp_tac. - masrp_tac. masrp_tac. masrp_tac. masrp_tac. masrp_tac. masrp_tac. masrp_tac. - masrp_tac. apply H8 in H1; inv_exists; simplify. repeat masrp_tac. auto. - repeat masrp_tac. - repeat masrp_tac. - repeat masrp_tac. - destruct (Pos.eq_dec (ram_mem r) s); subst; repeat masrp_tac. - destruct (Pos.eq_dec (ram_mem r) s); subst; repeat masrp_tac. - apply H9 in H17; auto. apply H9 in H17; auto. - Unshelve. eauto. -Qed. -#[local] Hint Resolve match_arrs_size_ram_preserved : mgen. - -Lemma match_arrs_size_ram_preserved2 : - forall m rs1 na na' ba ba' ram rs2, - exec_ram rs1 {| assoc_nonblocking := na; assoc_blocking := ba |} ram rs2 - {| assoc_nonblocking := na'; assoc_blocking := ba' |} -> - match_empty_size m na -> match_empty_size m ba -> - match_empty_size m na' /\ match_empty_size m ba'. -Proof. - intros. - remember ({| assoc_blocking := ba; assoc_nonblocking := na |}) as ar1. - remember ({| assoc_blocking := ba'; assoc_nonblocking := na' |}) as ar2. - assert (X1: na' = (assoc_nonblocking ar2)) by (rewrite Heqar2; auto). rewrite X1. - assert (X2: ba' = (assoc_blocking ar2)) by (rewrite Heqar2; auto). rewrite X2. - assert (X3: na = (assoc_nonblocking ar1)) by (rewrite Heqar1; auto). rewrite X3 in *. - assert (X4: ba = (assoc_blocking ar1)) by (rewrite Heqar1; auto). rewrite X4 in *. - eapply match_arrs_size_ram_preserved; mgen_crush. -Qed. -#[local] Hint Resolve match_arrs_size_ram_preserved2 : mgen. - -Lemma empty_stack_m : - forall m, empty_stack m = empty_stack' (mod_stk_len m) (mod_stk m). -Proof. unfold empty_stack, empty_stack'; mgen_crush. Qed. -#[local] Hint Rewrite empty_stack_m : mgen. - -Ltac clear_forall := - repeat match goal with - | H: context[forall _, _] |- _ => clear H - end. - -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; crush. - - symmetry. apply length_zero_iff_nil. auto. - - destruct l; crush. - rewrite list_repeat_cons. - crush. 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. - crush. - - rewrite <- (arr_wf a) in H. - apply list_combine_none. - assumption. -Qed. - -Lemma combine_none2 : - forall n a addr, - arr_length a = n -> - array_get_error addr (combine Verilog.merge_cell (arr_repeat None n) a) - = array_get_error addr a. -Proof. intros; auto using array_get_error_equal, combine_none. 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; crush. - - rewrite nth_error_nil in H0. - discriminate. - - destruct l2 eqn:EQl2. crush. - simpl in H. inv H. - destruct n; simpl in *. - inv 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; crush; auto. - - destruct l2 eqn:EQl2. crush. - simpl in H. inv H. - destruct n; simpl in *. - inv 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. - -Lemma match_get_arrs2 : - forall a i v l, - length a = l -> - list_combine merge_cell (list_set i (Some v) (list_repeat None l)) a = - list_combine merge_cell (list_repeat None l) (list_set i (Some v) a). -Proof. - induction a; crush; subst. - - destruct i. unfold list_repeat. unfold list_repeat'. auto. - unfold list_repeat. unfold list_repeat'. auto. - - destruct i. - rewrite list_repeat_cons. simplify. auto. - rewrite list_repeat_cons. simplify. f_equal. apply IHa. auto. -Qed. - -Lemma match_get_arrs : - forall addr i v x4 x x3, - x4 = arr_length x -> - x4 = arr_length x3 -> - array_get_error addr (combine merge_cell (array_set i (Some v) (arr_repeat None x4)) - (combine merge_cell x x3)) - = array_get_error addr (combine merge_cell (arr_repeat None x4) - (array_set i (Some v) (combine merge_cell x x3))). -Proof. - intros. apply array_get_error_equal. unfold combine. simplify. - destruct x; destruct x3; simplify. - apply match_get_arrs2. rewrite list_combine_length. subst. - rewrite H0. lia. -Qed. - -Lemma combine_array_set' : - forall a b i v, - length a = length b -> - list_combine merge_cell (list_set i (Some v) a) b = - list_set i (Some v) (list_combine merge_cell a b). -Proof. - induction a; simplify; crush. - - destruct i; crush. - - destruct i; destruct b; crush. - f_equal. apply IHa. auto. -Qed. - -Lemma combine_array_set : - forall a b i v addr, - arr_length a = arr_length b -> - array_get_error addr (combine merge_cell (array_set i (Some v) a) b) - = array_get_error addr (array_set i (Some v) (combine merge_cell a b)). -Proof. - intros. destruct a; destruct b. unfold array_set. simplify. - unfold array_get_error. simplify. f_equal. - apply combine_array_set'. crush. -Qed. - -Lemma array_get_combine' : - forall a b a' b' addr, - length a = length b -> - length a = length b' -> - length a = length a' -> - nth_error a addr = nth_error a' addr -> - nth_error b addr = nth_error b' addr -> - nth_error (list_combine merge_cell a b) addr = - nth_error (list_combine merge_cell a' b') addr. -Proof. - induction a; crush. - - destruct b; crush; destruct b'; crush; destruct a'; crush. - - destruct b; crush; destruct b'; crush; destruct a'; crush; - destruct addr; crush; apply IHa. -Qed. - -Lemma array_get_combine : - forall a b a' b' addr, - arr_length a = arr_length b -> - arr_length a = arr_length b' -> - arr_length a = arr_length a' -> - array_get_error addr a = array_get_error addr a' -> - array_get_error addr b = array_get_error addr b' -> - array_get_error addr (combine merge_cell a b) - = array_get_error addr (combine merge_cell a' b'). -Proof. - intros; unfold array_get_error, combine in *; destruct a; destruct b; - destruct a'; destruct b'; simplify; apply array_get_combine'; crush. -Qed. - -Lemma match_empty_size_exists_Some : - forall m rab s v, - match_empty_size m rab -> - rab ! s = Some v -> - exists v', (empty_stack m) ! s = Some v' /\ arr_length v = arr_length v'. -Proof. inversion 1; intros; repeat masrp_tac. Qed. - -Lemma match_empty_size_exists_None : - forall m rab s, - match_empty_size m rab -> - rab ! s = None -> - (empty_stack m) ! s = None. -Proof. inversion 1; intros; repeat masrp_tac. Qed. - -Lemma match_empty_size_exists_None' : - forall m rab s, - match_empty_size m rab -> - (empty_stack m) ! s = None -> - rab ! s = None. -Proof. inversion 1; intros; repeat masrp_tac. Qed. - -Lemma match_empty_size_exists_Some' : - forall m rab s v, - match_empty_size m rab -> - (empty_stack m) ! s = Some v -> - exists v', rab ! s = Some v' /\ arr_length v = arr_length v'. -Proof. inversion 1; intros; repeat masrp_tac. Qed. - -Lemma match_arrs_Some : - forall ra ra' s v, - match_arrs ra ra' -> - ra ! s = Some v -> - exists v', ra' ! s = Some v' - /\ (forall addr, array_get_error addr v = array_get_error addr v') - /\ arr_length v = arr_length v'. -Proof. inversion 1; intros; repeat masrp_tac. intros. rewrite H5. auto. Qed. - -Lemma match_arrs_None : - forall ra ra' s, - match_arrs ra ra' -> - ra ! s = None -> - ra' ! s = None. -Proof. inversion 1; intros; repeat masrp_tac. Qed. - -Ltac learn_next := - match goal with - | H: match_empty_size _ ?rab, H2: ?rab ! _ = Some _ |- _ => - let H3 := fresh "H" in - learn H2 as H3; eapply match_empty_size_exists_Some in H3; - eauto; inv_exists; simplify - | H: match_empty_size _ ?rab, H2: ?rab ! _ = None |- _ => - let H3 := fresh "H" in - learn H2 as H3; eapply match_empty_size_exists_None in H3; eauto - end. - -Ltac learn_empty := - match goal with - | H: match_empty_size _ _, H2: (empty_stack _) ! _ = Some _ |- _ => - let H3 := fresh "H" in - learn H as H3; eapply match_empty_size_exists_Some' in H3; - [| eassumption]; inv_exists; simplify - | H: match_arrs ?ar _, H2: ?ar ! _ = Some _ |- _ => - let H3 := fresh "H" in - learn H as H3; eapply match_arrs_Some in H3; - [| eassumption]; inv_exists; simplify - | H: match_empty_size _ _, H2: (empty_stack _) ! _ = None |- _ => - let H3 := fresh "H" in - learn H as H3; eapply match_empty_size_exists_None' in H3; - [| eassumption]; simplify - end. - -Lemma empty_set_none : - forall m ran rab s i v s0, - match_empty_size m ran -> - match_empty_size m rab -> - (arr_assocmap_set s i v ran) ! s0 = None -> - (arr_assocmap_set s i v rab) ! s0 = None. -Proof. - unfold arr_assocmap_set; inversion 1; subst; simplify. - destruct (Pos.eq_dec s s0); subst. - destruct ran ! s0 eqn:?. - rewrite AssocMap.gss in H4. inv H4. - learn_next. learn_empty. rewrite H6; auto. - destruct ran ! s eqn:?. rewrite AssocMap.gso in H4. - learn_next. learn_empty. rewrite H6. rewrite AssocMap.gso. - repeat match goal with - | H: Learnt _ |- _ => clear H - end. clear Heqo. clear H5. clear H6. - learn_next. repeat learn_empty. auto. auto. auto. - pose proof Heqo. learn_next; repeat learn_empty. - repeat match goal with - | H: Learnt _ |- _ => clear H - end. - pose proof H4. learn_next; repeat learn_empty. - rewrite H7. auto. -Qed. - -Ltac clear_learnt := - repeat match goal with - | H: Learnt _ |- _ => clear H - end. - -Lemma match_arrs_size_assoc : - forall a b, - match_arrs_size a b -> - match_arrs_size b a. -Proof. inversion 1; constructor; crush; split; apply H2. Qed. -#[local] Hint Resolve match_arrs_size_assoc : mgen. - -Lemma match_arrs_merge_set2 : - forall rab rab' ran ran' s m i v, - match_empty_size m rab -> - match_empty_size m ran -> - match_empty_size m rab' -> - match_empty_size m ran' -> - match_arrs rab rab' -> - match_arrs ran ran' -> - match_arrs (merge_arrs (arr_assocmap_set s i v ran) rab) - (merge_arrs (arr_assocmap_set s i v (empty_stack m)) - (merge_arrs ran' rab')). -Proof. - simplify. - constructor; intros. - unfold arr_assocmap_set in *. destruct (Pos.eq_dec s s0); subst. - destruct ran ! s0 eqn:?. unfold merge_arrs in *. rewrite AssocMap.gcombine in *; auto. - learn_next. repeat learn_empty. - econstructor. simplify. rewrite H6. rewrite AssocMap.gcombine by auto. - rewrite AssocMap.gss. simplify. setoid_rewrite H9. setoid_rewrite H7. simplify. - intros. rewrite AssocMap.gss in H5. setoid_rewrite H13 in H5. - simplify. pose proof (empty_arr m s0). inv H5. inv_exists. setoid_rewrite H5 in H6. inv H6. - unfold arr_repeat in H8. simplify. rewrite list_repeat_len in H8. rewrite list_repeat_len in H10. - rewrite match_get_arrs. crush. rewrite combine_none2. rewrite combine_array_set; try congruence. - apply array_get_error_each. rewrite combine_length; try congruence. - rewrite combine_length; try congruence. - apply array_get_combine; crush. - rewrite <- array_set_len. rewrite combine_length; crush. crush. crush. - setoid_rewrite H21 in H6; discriminate. rewrite combine_length. - rewrite <- array_set_len; crush. - unfold merge_arr in *. rewrite AssocMap.gss in H5. setoid_rewrite H13 in H5. - inv H5. rewrite combine_length. rewrite <- array_set_len; crush. - rewrite <- array_set_len; crush. - rewrite combine_length; crush. - destruct rab ! s0 eqn:?. learn_next. repeat learn_empty. - rewrite H11 in Heqo. discriminate. - unfold merge_arrs in H5. rewrite AssocMap.gcombine in H5; auto. rewrite Heqo in H5. - rewrite Heqo0 in H5. crush. - - destruct ran ! s eqn:?. - learn_next. repeat learn_empty. rewrite H6. - unfold merge_arrs in *. rewrite AssocMap.gcombine in H5; auto. - rewrite AssocMap.gcombine; auto. rewrite AssocMap.gso in H5; auto. - rewrite AssocMap.gso; auto. - destruct ran ! s0 eqn:?. - learn_next. - repeat match goal with - | H: Learnt _ |- _ => clear H - end. - repeat learn_empty. - repeat match goal with - | H: Learnt _ |- _ => clear H - end. - rewrite AssocMap.gcombine; auto. setoid_rewrite Heqo0 in H5. setoid_rewrite H29 in H5. - simplify. - pose proof (empty_arr m s0). inv H5. inv_exists. rewrite H5 in H21. inv H21. - econstructor. simplify. setoid_rewrite H23. rewrite H25. setoid_rewrite H5. - simplify. intros. rewrite combine_none2. apply array_get_combine; solve [crush]. - crush. rewrite list_combine_length. rewrite (arr_wf x5). rewrite (arr_wf x6). - rewrite <- H26. rewrite <- H28. rewrite list_repeat_len. lia. rewrite list_combine_length. - rewrite (arr_wf a). rewrite (arr_wf x7). rewrite combine_length. rewrite arr_repeat_length. - rewrite H24. rewrite <- H32. rewrite list_repeat_len. lia. - rewrite arr_repeat_length. rewrite combine_length. rewrite <- H26. symmetry. apply list_repeat_len. - congruence. - rewrite H37 in H21; discriminate. - repeat match goal with - | H: Learnt _ |- _ => clear H - end. eapply match_empty_size_exists_None in H0; eauto. - clear H6. repeat learn_empty. setoid_rewrite Heqo0 in H5. - setoid_rewrite H29 in H5. discriminate. - pose proof (match_arrs_merge ran ran' rab rab'). - eapply match_empty_size_match in H; [|apply H0]. - apply H6 in H; auto. inv H. apply H7 in H5. inv_exists. simplify. - learn_next. rewrite H9. econstructor. simplify. - apply merge_arr_empty''; mgen_crush. - auto. auto. - unfold merge_arrs in *. rewrite AssocMap.gcombine in H5; auto. rewrite AssocMap.gcombine; auto. - destruct (arr_assocmap_set s i v ran) ! s0 eqn:?; destruct rab ! s0 eqn:?; crush. - learn_next. repeat learn_empty. - repeat match goal with - | H: Learnt _ |- _ => clear H - end. - erewrite empty_set_none. rewrite AssocMap.gcombine; auto. - simplify. rewrite H7. rewrite H8. auto. apply H0. mgen_crush. auto. -Qed. - -Definition all_match_empty_size m ar := - match_empty_size m (assoc_nonblocking ar) /\ match_empty_size m (assoc_blocking ar). -#[local] Hint Unfold all_match_empty_size : mgen. - -Definition match_module_to_ram m ram := - ram_size ram = mod_stk_len m /\ ram_mem ram = mod_stk m. -#[local] Hint Unfold match_module_to_ram : mgen. - -Lemma zip_range_forall_le : - forall A (l: list A) n, Forall (Pos.le n) (map snd (zip_range n l)). -Proof. - induction l; crush; constructor; [lia|]. - assert (forall n x, n+1 <= x -> n <= x) by lia. - apply Forall_forall. intros. apply H. generalize dependent x. - apply Forall_forall. apply IHl. -Qed. - -Lemma transf_code_fold_correct: - forall l m state ram d' c' n, - fold_right (transf_maps state ram) (mod_datapath m, mod_controllogic m) l = (d', c') -> - Forall (fun x => x < n) (map fst l) -> - Forall (Pos.le n) (map snd l) -> - list_norepet (map fst l) -> - list_norepet (map snd l) -> - (forall p i c_s rs1 ar1 rs2 ar2 trs1 tar1 d_s, - i < n -> - all_match_empty_size m ar1 -> - all_match_empty_size m tar1 -> - match_module_to_ram m ram -> - (mod_datapath m)!i = Some d_s -> - (mod_controllogic m)!i = Some c_s -> - match_reg_assocs p rs1 trs1 -> - match_arr_assocs ar1 tar1 -> - max_reg_module m < p -> - exec_all d_s c_s rs1 ar1 rs2 ar2 -> - exists d_s' c_s' trs2 tar2, - d'!i = Some d_s' /\ c'!i = Some c_s' - /\ exec_all_ram ram d_s' c_s' trs1 tar1 trs2 tar2 - /\ match_reg_assocs p (merge_reg_assocs rs2) (merge_reg_assocs trs2) - /\ match_arr_assocs (merge_arr_assocs (ram_mem ram) (ram_size ram) ar2) - (merge_arr_assocs (ram_mem ram) (ram_size ram) tar2)). -Proof. - induction l as [| a l IHl]; simplify. - - match goal with - | H: (_, _) = (_, _) |- _ => inv H - end; - unfold exec_all in *; repeat inv_exists; simplify. - exploit match_states_same; - try match goal with - | H: stmnt_runp _ _ _ ?c _ _, H2: (mod_controllogic _) ! _ = Some ?c |- _ => apply H - end; eauto; mgen_crush; - try match goal with - | H: (mod_controllogic _) ! _ = Some ?c |- _ => - apply max_reg_stmnt_le_stmnt_tree in H; unfold max_reg_module in *; lia - end; intros; - exploit match_states_same; - try match goal with - | H: stmnt_runp _ _ _ ?c _ _, H2: (mod_datapath _) ! _ = Some ?c |- _ => apply H - end; eauto; mgen_crush; - try match goal with - | H: (mod_datapath _) ! _ = Some ?c |- _ => - apply max_reg_stmnt_le_stmnt_tree in H; unfold max_reg_module in *; lia - end; intros; - repeat match goal with - | |- exists _, _ => eexists - end; simplify; eauto; - unfold exec_all_ram; - repeat match goal with - | |- exists _, _ => eexists - end; simplify; eauto. - constructor. admit. - Abort. - -Lemma empty_stack_transf : forall m, empty_stack (transf_module m) = empty_stack m. -Proof. unfold empty_stack, transf_module; intros; repeat destruct_match; crush. Qed. - -Definition alt_unchanged (d : AssocMap.t stmnt) (c: AssocMap.t stmnt) d' c' i := - d ! i = d' ! i /\ c ! i = c' ! i. - -Definition alt_store ram d (c : AssocMap.t stmnt) d' c' i := - exists e1 e2, - d' ! i = Some (Vseq (Vnonblock (Vvar (ram_u_en ram)) (Vunop Vnot (Vvar (ram_u_en ram)))) - (Vseq (Vnonblock (Vvar (ram_wr_en ram)) (Vlit (ZToValue 1))) - (Vseq (Vnonblock (Vvar (ram_d_in ram)) e2) - (Vnonblock (Vvar (ram_addr ram)) e1)))) - /\ c' ! i = c ! i - /\ d ! i = Some (Vnonblock (Vvari (ram_mem ram) e1) e2). - -Definition alt_load state ram d (c : AssocMap.t stmnt) d' c' i n := - exists ns e1 e2, - d' ! i = Some (Vseq (Vnonblock (Vvar (ram_u_en ram)) (Vunop Vnot (Vvar (ram_u_en ram)))) - (Vseq (Vnonblock (Vvar (ram_wr_en ram)) (Vlit (ZToValue 0))) - (Vnonblock (Vvar (ram_addr ram)) e2))) - /\ d' ! n = Some (Vnonblock (Vvar e1) (Vvar (ram_d_out ram))) - /\ c' ! i = Some (Vnonblock (Vvar state) (Vlit (posToValue n))) - /\ c' ! n = Some (Vnonblock (Vvar state) ns) - /\ c ! i = Some (Vnonblock (Vvar state) ns) - /\ d ! i = Some (Vnonblock (Vvar e1) (Vvari (ram_mem ram) e2)) - /\ e1 < state - /\ max_reg_expr e2 < state - /\ max_reg_expr ns < state - /\ (Z.pos n <= Int.max_unsigned)%Z. - -Definition alternatives state ram d c d' c' i n := - alt_unchanged d c d' c' i - \/ alt_store ram d c d' c' i - \/ alt_load state ram d c d' c' i n. - -Lemma transf_alternatives : - forall ram n d c state i d' c', - transf_maps state ram (i, n) (d, c) = (d', c') -> - i <> n -> - alternatives state ram d c d' c' i n. -Proof. - intros. unfold transf_maps in *. - repeat destruct_match; match goal with - | H: (_, _) = (_, _) |- _ => inv H - end; try solve [left; econstructor; crush]; simplify; - repeat match goal with - | H: (_ =? _) = true |- _ => apply Peqb_true_eq in H; subst - end; unfold alternatives; right; - match goal with - | H: context[Vnonblock (Vvari _ _) _] |- _ => left - | _ => right - end; repeat econstructor; simplify; - repeat match goal with - | |- ( _ # ?s <- _ ) ! ?s = Some _ => apply AssocMap.gss - | |- ( _ # ?s <- _ ) ! ?r = Some _ => rewrite AssocMap.gso by lia - | |- _ = None => apply max_index_2; lia - | H: context[_ apply Pos.ltb_lt in H - end; auto. -Qed. - -Lemma transf_alternatives_neq : - forall state ram a n' n d'' c'' d' c' i d c, - transf_maps state ram (a, n) (d, c) = (d', c') -> - a <> i -> n' <> n -> i <> n' -> a <> n' -> - i <> n -> a <> n -> - alternatives state ram d'' c'' d c i n' -> - alternatives state ram d'' c'' d' c' i n'. -Proof. - unfold alternatives, alt_unchanged, alt_store, alt_load, transf_maps; intros; - repeat match goal with H: _ \/ _ |- _ => inv H | H: _ /\ _ |- _ => destruct H end; - [left | right; left | right; right]; - repeat inv_exists; simplify; - repeat destruct_match; - repeat match goal with - | H: (_, _) = (_, _) |- _ => inv H - | |- exists _, _ => econstructor - end; repeat split; repeat rewrite AssocMap.gso by lia; eauto; lia. -Qed. - -Lemma transf_alternatives_neq2 : - forall state ram a n' n d'' c'' d' c' i d c, - transf_maps state ram (a, n) (d', c') = (d, c) -> - a <> i -> n' <> n -> i <> n' -> a <> n' -> i <> n -> - alternatives state ram d c d'' c'' i n' -> - alternatives state ram d' c' d'' c'' i n'. -Proof. - unfold alternatives, alt_unchanged, alt_store, alt_load, transf_maps; intros; - repeat match goal with H: _ \/ _ |- _ => inv H | H: _ /\ _ |- _ => destruct H end; - [left | right; left | right; right]; - repeat inv_exists; simplify; - repeat destruct_match; - repeat match goal with - | H: (_, _) = (_, _) |- _ => inv H - | |- exists _, _ => econstructor - end; repeat split; repeat rewrite AssocMap.gso in * by lia; eauto; lia. -Qed. - -Lemma transf_alt_unchanged_neq : - forall i c'' d'' d c d' c', - alt_unchanged d' c' d'' c'' i -> - d' ! i = d ! i -> - c' ! i = c ! i -> - alt_unchanged d c d'' c'' i. -Proof. unfold alt_unchanged; simplify; congruence. Qed. - -Lemma transf_maps_neq : - forall state ram d c i n d' c' i' n' va vb vc vd, - transf_maps state ram (i, n) (d, c) = (d', c') -> - d ! i' = va -> d ! n' = vb -> - c ! i' = vc -> c ! n' = vd -> - i <> i' -> i <> n' -> n <> i' -> n <> n' -> - d' ! i' = va /\ d' ! n' = vb /\ c' ! i' = vc /\ c' ! n' = vd. -Proof. - unfold transf_maps; intros; repeat destruct_match; simplify; - repeat match goal with - | H: (_, _) = (_, _) |- _ => inv H - | H: (_ =? _) = true |- _ => apply Peqb_true_eq in H; subst - | |- context[( _ # ?s <- _ ) ! ?r] => rewrite AssocMap.gso by lia - end; crush. -Qed. - -Lemma alternatives_different_map : - forall l state ram d c d'' c'' d' c' n i p, - i <= p -> n > p -> - Forall (Pos.lt p) (map snd l) -> - Forall (Pos.ge p) (map fst l) -> - ~ In n (map snd l) -> - ~ In i (map fst l) -> - fold_right (transf_maps state ram) (d, c) l = (d', c') -> - alternatives state ram d' c' d'' c'' i n -> - alternatives state ram d c d'' c'' i n. -Proof. - Opaque transf_maps. - induction l; intros. - - crush. - - simplify; repeat match goal with - | H: context[_ :: _] |- _ => inv H - | H: transf_maps _ _ _ (fold_right (transf_maps ?s ?r) (?d, ?c) ?l) = (_, _) |- _ => - let X := fresh "X" in - remember (fold_right (transf_maps s r) (d, c) l) as X - | X: _ * _ |- _ => destruct X - | H: (_, _) = _ |- _ => symmetry in H - end; simplify. - remember p0 as i'. symmetry in Heqi'. subst. - remember p1 as n'. symmetry in Heqn'. subst. - assert (i <> n') by lia. - assert (n <> i') by lia. - assert (n <> n') by lia. - assert (i <> i') by lia. - eapply IHl; eauto. - eapply transf_alternatives_neq2; eauto; try lia. -Qed. - -Lemma transf_fold_alternatives : - forall l state ram d c d' c' i n d_s c_s, - fold_right (transf_maps state ram) (d, c) l = (d', c') -> - Pos.max (max_pc c) (max_pc d) < n -> - Forall (Pos.lt (Pos.max (max_pc c) (max_pc d))) (map snd l) -> - Forall (Pos.ge (Pos.max (max_pc c) (max_pc d))) (map fst l) -> - list_norepet (map fst l) -> - list_norepet (map snd l) -> - In (i, n) l -> - d ! i = Some d_s -> - c ! i = Some c_s -> - alternatives state ram d c d' c' i n. -Proof. - Opaque transf_maps. - induction l; crush; []. - repeat match goal with - | H: context[_ :: _] |- _ => inv H - | H: transf_maps _ _ _ (fold_right (transf_maps ?s ?r) (?d, ?c) ?l) = (_, _) |- _ => - let X := fresh "X" in - remember (fold_right (transf_maps s r) (d, c) l) as X - | X: _ * _ |- _ => destruct X - | H: (_, _) = _ |- _ => symmetry in H - end. - inv H5. inv H1. simplify. - eapply alternatives_different_map; eauto. - simplify; lia. simplify; lia. apply transf_alternatives; auto. lia. - simplify. - assert (X: In i (map fst l)). { replace i with (fst (i, n)) by auto. apply in_map; auto. } - assert (X2: In n (map snd l)). { replace n with (snd (i, n)) by auto. apply in_map; auto. } - assert (X3: n <> p0). { destruct (Pos.eq_dec n p0); subst; crush. } - assert (X4: i <> p). { destruct (Pos.eq_dec i p); subst; crush. } - eapply transf_alternatives_neq; eauto; apply max_index in H7; lia. - Transparent transf_maps. -Qed. - -Lemma zip_range_inv : - forall A (l: list A) i n, - In i l -> - exists n', In (i, n') (zip_range n l) /\ n' >= n. -Proof. - induction l; crush. - inv H. econstructor. - split. left. eauto. lia. - eapply IHl in H0. inv H0. inv H. - econstructor. split. right. apply H0. lia. -Qed. - -Lemma zip_range_not_in_fst : - forall A (l: list A) a n, ~ In a l -> ~ In a (map fst (zip_range n l)). -Proof. unfold not; induction l; crush; inv H0; firstorder. Qed. - -Lemma zip_range_no_repet_fst : - forall A (l: list A) a, list_norepet l -> list_norepet (map fst (zip_range a l)). -Proof. - induction l; simplify; constructor; inv H; firstorder; - eapply zip_range_not_in_fst; auto. -Qed. - -Lemma transf_code_alternatives : - forall state ram d c d' c' i d_s c_s, - transf_code state ram d c = (d', c') -> - d ! i = Some d_s -> - c ! i = Some c_s -> - exists n, alternatives state ram d c d' c' i n. -Proof. - unfold transf_code; - intros. - pose proof H0 as X. - apply PTree.elements_correct in X. assert (In i (map fst (PTree.elements d))). - { replace i with (fst (i, d_s)) by auto. apply in_map. auto. } - exploit zip_range_inv. apply H2. intros. inv H3. simplify. - instantiate (1 := (Pos.max (max_pc c) (max_pc d) + 1)) in H3. - exists x. - eapply transf_fold_alternatives; - eauto using forall_gt, PTree.elements_keys_norepet, max_index. lia. - assert (Forall (Pos.le (Pos.max (max_pc c) (max_pc d) + 1)) - (map snd (zip_range (Pos.max (max_pc c) (max_pc d) + 1) - (map fst (PTree.elements d))))) by apply zip_range_forall_le. - apply Forall_forall; intros. eapply Forall_forall in H4; eauto. lia. - rewrite zip_range_fst_idem. apply Forall_forall; intros. - apply AssocMapExt.elements_iff in H4. inv H4. apply max_index in H6. lia. - eapply zip_range_no_repet_fst. apply PTree.elements_keys_norepet. - eapply zip_range_snd_no_repet. -Qed. - -Lemma max_reg_stmnt_not_modified : - forall s f rs ar rs' ar', - stmnt_runp f rs ar s rs' ar' -> - forall r, - max_reg_stmnt s < r -> - (assoc_blocking rs) ! r = (assoc_blocking rs') ! r. -Proof. - induction 1; crush; - try solve [repeat destruct_match; apply IHstmnt_runp; try lia; auto]. - assert (X: (assoc_blocking asr1) ! r = (assoc_blocking asr2) ! r) by (apply IHstmnt_runp2; lia). - assert (X2: (assoc_blocking asr0) ! r = (assoc_blocking asr1) ! r) by (apply IHstmnt_runp1; lia). - congruence. - inv H. simplify. rewrite AssocMap.gso by lia; auto. -Qed. - -Lemma max_reg_stmnt_not_modified_nb : - forall s f rs ar rs' ar', - stmnt_runp f rs ar s rs' ar' -> - forall r, - max_reg_stmnt s < r -> - (assoc_nonblocking rs) ! r = (assoc_nonblocking rs') ! r. -Proof. - induction 1; crush; - try solve [repeat destruct_match; apply IHstmnt_runp; try lia; auto]. - assert (X: (assoc_nonblocking asr1) ! r = (assoc_nonblocking asr2) ! r) by (apply IHstmnt_runp2; lia). - assert (X2: (assoc_nonblocking asr0) ! r = (assoc_nonblocking asr1) ! r) by (apply IHstmnt_runp1; lia). - congruence. - inv H. simplify. rewrite AssocMap.gso by lia; auto. -Qed. - -Lemma int_eq_not_changed : - forall ar ar' r r2 b, - Int.eq (ar # r) (ar # r2) = b -> - ar ! r = ar' ! r -> - ar ! r2 = ar' ! r2 -> - Int.eq (ar' # r) (ar' # r2) = b. -Proof. - unfold find_assocmap, AssocMapExt.get_default. intros. - rewrite <- H0. rewrite <- H1. auto. -Qed. - -Lemma merge_find_assocmap : - forall ran rab x, - ran ! x = None -> - (merge_regs ran rab) # x = rab # x. -Proof. - unfold merge_regs, find_assocmap, AssocMapExt.get_default. - intros. destruct (rab ! x) eqn:?. - erewrite AssocMapExt.merge_correct_2; eauto. - erewrite AssocMapExt.merge_correct_3; eauto. -Qed. - -Lemma max_reg_module_controllogic_gt : - forall m i v p, - (mod_controllogic m) ! i = Some v -> - max_reg_module m < p -> - max_reg_stmnt v < p. -Proof. - intros. unfold max_reg_module in *. - apply max_reg_stmnt_le_stmnt_tree in H. lia. -Qed. - -Lemma max_reg_module_datapath_gt : - forall m i v p, - (mod_datapath m) ! i = Some v -> - max_reg_module m < p -> - max_reg_stmnt v < p. -Proof. - intros. unfold max_reg_module in *. - apply max_reg_stmnt_le_stmnt_tree in H. lia. -Qed. - -Lemma merge_arr_empty2 : - forall m ar ar', - match_empty_size m ar' -> - match_arrs ar ar' -> - match_arrs ar (merge_arrs (empty_stack m) ar'). -Proof. - inversion 1; subst; inversion 1; subst. - econstructor; intros. apply H4 in H6; inv_exists. simplify. - eapply merge_arr_empty'' in H6; eauto. - (* apply H5 in H6. pose proof H6. apply H2 in H7. *) -(* unfold merge_arrs. rewrite AssocMap.gcombine; auto. setoid_rewrite H6. *) -(* setoid_rewrite H7. auto. *) -(* Qed. *) Admitted. -#[local] Hint Resolve merge_arr_empty2 : mgen. - -Lemma find_assocmap_gso : - forall ar x y v, x <> y -> (ar # y <- v) # x = ar # x. -Proof. - unfold find_assocmap, AssocMapExt.get_default; intros; rewrite AssocMap.gso; auto. -Qed. - -Lemma find_assocmap_gss : - forall ar x v, (ar # x <- v) # x = v. -Proof. - unfold find_assocmap, AssocMapExt.get_default; intros; rewrite AssocMap.gss; auto. -Qed. - -Lemma expr_lt_max_module_datapath : - forall m x, - max_reg_stmnt x <= max_stmnt_tree (mod_datapath m) -> - max_reg_stmnt x < max_reg_module m + 1. -Proof. unfold max_reg_module. lia. Qed. - -Lemma expr_lt_max_module_controllogic : - forall m x, - max_reg_stmnt x <= max_stmnt_tree (mod_controllogic m) -> - max_reg_stmnt x < max_reg_module m + 1. -Proof. unfold max_reg_module. lia. Qed. - -Lemma int_eq_not : - forall x y, Int.eq x y = true -> Int.eq x (Int.not y) = false. -Proof. - intros. pose proof (Int.eq_spec x y). rewrite H in H0. subst. - apply int_eq_not_false. -Qed. - -Lemma match_assocmaps_gt2 : - forall (p s : positive) (ra ra' : assocmap) (v : value), - p <= s -> match_assocmaps p ra ra' -> match_assocmaps p (ra # s <- v) ra'. -Proof. - intros; inv H0; constructor; intros. - destruct (Pos.eq_dec r s); subst. lia. - rewrite AssocMap.gso by lia. auto. -Qed. - -Lemma match_assocmaps_switch_neq : - forall p ra ra' r v' s v, - match_assocmaps p ra ((ra' # r <- v') # s <- v) -> - s <> r -> - match_assocmaps p ra ((ra' # s <- v) # r <- v'). -Proof. - inversion 1; constructor; simplify. - destruct (Pos.eq_dec r0 s); destruct (Pos.eq_dec r0 r); subst; try lia. - rewrite AssocMap.gso by lia. specialize (H0 s). apply H0 in H5. - rewrite AssocMap.gss in H5. rewrite AssocMap.gss. auto. - rewrite AssocMap.gss. apply H0 in H5. rewrite AssocMap.gso in H5 by lia. - rewrite AssocMap.gss in H5. auto. - repeat rewrite AssocMap.gso by lia. - apply H0 in H5. repeat rewrite AssocMap.gso in H5 by lia. - auto. -Qed. - -Lemma match_assocmaps_duplicate : - forall p ra ra' v' s v, - match_assocmaps p ra (ra' # s <- v) -> - match_assocmaps p ra ((ra' # s <- v') # s <- v). -Proof. - inversion 1; constructor; simplify. - destruct (Pos.eq_dec r s); subst. - rewrite AssocMap.gss. apply H0 in H4. rewrite AssocMap.gss in H4. auto. - repeat rewrite AssocMap.gso by lia. apply H0 in H4. rewrite AssocMap.gso in H4 by lia. - auto. -Qed. - -Ltac simplify := intros; unfold_constants; cbn in *; - repeat (nicify_hypotheses; nicify_goals; kill_bools; substpp); - cbn in *. - -Ltac simplify_val := repeat (simplify; unfold uvalueToZ, valueToPtr, Ptrofs.of_int, valueToInt, intToValue, - ptrToValue in *). - -Ltac crush_val := simplify_val; try discriminate; try congruence; try lia; liapp; try assumption. - -Ltac crush := simplify; try discriminate; try congruence; try lia; liapp; - try assumption; auto. - -Ltac mgen_crush_local := crush; eauto with mgen. - -Lemma translation_correct : - forall m asr nasa1 basa1 nasr1 basr1 basr2 nasr2 nasa2 basa2 nasr3 basr3 - nasa3 basa3 asr'0 asa'0 res' st tge pstval sf asa ctrl data f, - asr ! (mod_reset m) = Some (ZToValue 0) -> - asr ! (mod_finish m) = Some (ZToValue 0) -> - asr ! (mod_st m) = Some (posToValue st) -> - (mod_controllogic m) ! st = Some ctrl -> - (mod_datapath m) ! st = Some data -> - stmnt_runp f {| assoc_blocking := asr; assoc_nonblocking := empty_assocmap |} - {| assoc_blocking := asa; assoc_nonblocking := empty_stack m |} ctrl - {| assoc_blocking := basr1; assoc_nonblocking := nasr1 |} - {| assoc_blocking := basa1; assoc_nonblocking := nasa1 |} -> - basr1 ! (mod_st m) = Some (posToValue st) -> - stmnt_runp f {| assoc_blocking := basr1; assoc_nonblocking := nasr1 |} - {| assoc_blocking := basa1; assoc_nonblocking := nasa1 |} data - {| assoc_blocking := basr2; assoc_nonblocking := nasr2 |} - {| assoc_blocking := basa2; assoc_nonblocking := nasa2 |} -> - exec_ram {| assoc_blocking := merge_regs nasr2 basr2; assoc_nonblocking := empty_assocmap |} - {| assoc_blocking := merge_arrs nasa2 basa2; assoc_nonblocking := empty_stack m |} None - {| assoc_blocking := basr3; assoc_nonblocking := nasr3 |} - {| assoc_blocking := basa3; assoc_nonblocking := nasa3 |} -> - (merge_regs nasr3 basr3) ! (mod_st m) = Some (posToValue pstval) -> - (Z.pos pstval <= 4294967295)%Z -> - match_states (State sf m st asr asa) (State res' (transf_module m) st asr'0 asa'0) -> - mod_ram m = None -> - exists R2 : state, - Smallstep.plus step tge (State res' (transf_module m) st asr'0 asa'0) Events.E0 R2 /\ - match_states (State sf m pstval (merge_regs nasr3 basr3) (merge_arrs nasa3 basa3)) R2. -Proof. - Ltac tac0 := - repeat match goal with - | H: match_reg_assocs _ _ _ |- _ => inv H - | H: match_arr_assocs _ _ |- _ => inv H - end. - intros. - repeat match goal with - | H: match_states _ _ |- _ => inv H - | H: context[exec_ram] |- _ => inv H - | H: mod_ram _ = None |- _ => - let H2 := fresh "TRANSF" in learn H as H2; apply transf_module_code in H2 - end. - eapply transf_code_alternatives in TRANSF; eauto; simplify; unfold alternatives in *. - repeat match goal with H: _ \/ _ |- _ => inv H end. - - unfold alt_unchanged in *; simplify. - assert (MATCH_SIZE1: match_empty_size m nasa1 /\ match_empty_size m basa1). - { eapply match_arrs_size_stmnt_preserved2; eauto. unfold match_empty_size; eauto with mgen. } - assert (MATCH_SIZE2: match_empty_size m nasa2 /\ match_empty_size m basa2). - { eapply match_arrs_size_stmnt_preserved2; mgen_crush. } simplify. - assert (MATCH_ARR3: match_arrs_size nasa2 basa2) by eauto with mgen. - exploit match_states_same; try solve [apply H4 | eapply max_stmnt_lt_module; eauto - | econstructor; eauto with mgen]; - intros; repeat inv_exists; simplify; tac0. - exploit match_states_same; try solve [eapply H6 | eapply max_stmnt_lt_module; eauto - | econstructor; eauto with mgen]; - intros; repeat inv_exists; simplify; tac0. - assert (MATCH_SIZE1': match_empty_size m ran'0 /\ match_empty_size m rab'0). - { eapply match_arrs_size_stmnt_preserved2; eauto. unfold match_empty_size; eauto with mgen. - rewrite empty_stack_transf; eauto with mgen. } - assert (MATCH_SIZE2': match_empty_size m ran'2 /\ match_empty_size m rab'2). - { eapply match_arrs_size_stmnt_preserved2; mgen_crush. } simplify. - assert (MATCH_ARR3': match_arrs_size ran'2 rab'2) by eauto with mgen. - do 2 econstructor. apply Smallstep.plus_one. econstructor. - eauto with mgen. eauto with mgen. eauto with mgen. - rewrite <- H12. eassumption. rewrite <- H7. eassumption. - eauto. eauto with mgen. eauto. - rewrite empty_stack_transf. unfold transf_module; repeat destruct_match; try discriminate. - econstructor. simplify. - unfold disable_ram in *. unfold transf_module in DISABLE_RAM. - repeat destruct_match; try discriminate; []. simplify. - pose proof H17 as R1. apply max_reg_stmnt_not_modified with (r := max_reg_module m + 2) in R1. - pose proof H17 as R2. eapply max_reg_stmnt_not_modified_nb with (r := max_reg_module m + 2) in R2. - pose proof H18 as R3. eapply max_reg_stmnt_not_modified with (r := max_reg_module m + 2) in R3. - pose proof H18 as R4. eapply max_reg_stmnt_not_modified_nb with (r := max_reg_module m + 2) in R4. - pose proof H17 as R1'. apply max_reg_stmnt_not_modified with (r := max_reg_module m + 6) in R1'. - pose proof H17 as R2'. eapply max_reg_stmnt_not_modified_nb with (r := max_reg_module m + 6) in R2'. - pose proof H18 as R3'. eapply max_reg_stmnt_not_modified with (r := max_reg_module m + 6) in R3'. - pose proof H18 as R4'. eapply max_reg_stmnt_not_modified_nb with (r := max_reg_module m + 6) in R4'. - simplify. - pose proof DISABLE_RAM as DISABLE_RAM1. - eapply int_eq_not_changed with (ar' := rab') in DISABLE_RAM; try congruence. - eapply int_eq_not_changed with (ar' := rab'1) in DISABLE_RAM; try congruence. - rewrite AssocMap.gempty in R2. rewrite <- R2 in R4. - rewrite AssocMap.gempty in R2'. rewrite <- R2' in R4'. - eapply int_eq_not_changed in DISABLE_RAM; auto. repeat (rewrite merge_find_assocmap; try congruence). - eapply max_reg_module_datapath_gt; eauto; remember (max_reg_module m); lia. - eapply max_reg_module_datapath_gt; eauto; remember (max_reg_module m); lia. - eapply max_reg_module_controllogic_gt; eauto; remember (max_reg_module m); lia. - eapply max_reg_module_controllogic_gt; eauto; remember (max_reg_module m); lia. - eapply max_reg_module_datapath_gt; eauto; remember (max_reg_module m); lia. - eapply max_reg_module_datapath_gt; eauto; remember (max_reg_module m); lia. - eapply max_reg_module_controllogic_gt; eauto; remember (max_reg_module m); lia. - eapply max_reg_module_controllogic_gt; eauto; remember (max_reg_module m); lia. - auto. auto. eauto with mgen. auto. - econstructor; mgen_crush_local. apply merge_arr_empty; mgen_crush_local. - unfold disable_ram in *. unfold transf_module in DISABLE_RAM. - repeat destruct_match; crush. unfold transf_module in Heqo; repeat destruct_match; crush. - pose proof H17 as R1. apply max_reg_stmnt_not_modified with (r := max_reg_module m + 2) in R1. - pose proof H17 as R2. eapply max_reg_stmnt_not_modified_nb with (r := max_reg_module m + 2) in R2. - pose proof H18 as R3. eapply max_reg_stmnt_not_modified with (r := max_reg_module m + 2) in R3. - pose proof H18 as R4. eapply max_reg_stmnt_not_modified_nb with (r := max_reg_module m + 2) in R4. - pose proof H17 as R1'. apply max_reg_stmnt_not_modified with (r := max_reg_module m + 6) in R1'. - pose proof H17 as R2'. eapply max_reg_stmnt_not_modified_nb with (r := max_reg_module m + 6) in R2'. - pose proof H18 as R3'. eapply max_reg_stmnt_not_modified with (r := max_reg_module m + 6) in R3'. - pose proof H18 as R4'. eapply max_reg_stmnt_not_modified_nb with (r := max_reg_module m + 6) in R4'. - simplify. - pose proof DISABLE_RAM as DISABLE_RAM1. - eapply int_eq_not_changed with (ar' := rab') in DISABLE_RAM; try congruence. - eapply int_eq_not_changed with (ar' := rab'1) in DISABLE_RAM; try congruence. - rewrite AssocMap.gempty in R2. rewrite <- R2 in R4. - rewrite AssocMap.gempty in R2'. rewrite <- R2' in R4'. - eapply int_eq_not_changed in DISABLE_RAM; auto. repeat (rewrite merge_find_assocmap; try congruence). - unfold empty_assocmap; apply AssocMap.gempty. - unfold empty_assocmap; apply AssocMap.gempty. - eapply max_reg_module_datapath_gt; eauto; remember (max_reg_module m); lia. - eapply max_reg_module_datapath_gt; eauto; remember (max_reg_module m); lia. - eapply max_reg_module_controllogic_gt; eauto; remember (max_reg_module m); lia. - eapply max_reg_module_controllogic_gt; eauto; remember (max_reg_module m); lia. - eapply max_reg_module_datapath_gt; eauto; remember (max_reg_module m); lia. - eapply max_reg_module_datapath_gt; eauto; remember (max_reg_module m); lia. - eapply max_reg_module_controllogic_gt; eauto; remember (max_reg_module m); lia. - eapply max_reg_module_controllogic_gt; eauto; remember (max_reg_module m); lia. - - unfold alt_store in *; simplify. inv H6. inv H19. inv H19. simplify. - exploit match_states_same; try solve [eapply H4 | eapply max_stmnt_lt_module; eauto - | econstructor; eauto with mgen]; - intros; repeat inv_exists; simplify; tac0. - do 2 econstructor. apply Smallstep.plus_one. econstructor. solve [eauto with mgen]. solve [eauto with mgen]. - solve [eauto with mgen]. - rewrite H7. eassumption. eassumption. eassumption. solve [eauto with mgen]. - econstructor. econstructor. econstructor. econstructor. econstructor. - auto. auto. auto. econstructor. econstructor. econstructor. - econstructor. econstructor. econstructor. econstructor. - eapply expr_runp_matches2. eassumption. 2: { cbn. eassumption. } - pose proof H3 as X. apply max_reg_stmnt_le_stmnt_tree in X. - apply expr_lt_max_module_datapath in X; simplify; remember (max_reg_module m); lia. - auto. - econstructor. econstructor. eapply expr_runp_matches2; eauto. - pose proof H3 as X. apply max_reg_stmnt_le_stmnt_tree in X. - apply expr_lt_max_module_datapath in X. - remember (max_reg_module m); simplify; lia. - - rewrite empty_stack_transf. - unfold transf_module; repeat destruct_match; try discriminate; simplify; []. - eapply exec_ram_Some_write. - 3: { - simplify. unfold merge_regs. repeat rewrite AssocMapExt.merge_add_assoc. - repeat rewrite find_assocmap_gso by lia. - pose proof H12 as X. - eapply max_reg_stmnt_not_modified_nb with (r := (max_reg_module m + 2)) in X. - rewrite AssocMap.gempty in X. - apply merge_find_assocmap. repeat rewrite AssocMap.gso by lia. auto. - apply max_reg_stmnt_le_stmnt_tree in H2. - apply expr_lt_max_module_controllogic in H2. lia. - } - 3: { - simplify. unfold merge_regs. erewrite merge_get_default4. eauto. - repeat rewrite AssocMap.gso by lia. apply AssocMap.gss. - } - { unfold disable_ram in *. unfold transf_module in DISABLE_RAM; - repeat destruct_match; try discriminate. - simplify. - pose proof H12 as X2. - pose proof H12 as X4. - apply max_reg_stmnt_not_modified with (r := max_reg_module m + 2) in X2. - apply max_reg_stmnt_not_modified with (r := max_reg_module m + 6) in X4. - assert (forall ar ar' x, ar ! x = ar' ! x -> ar # x = ar' # x). - { intros. unfold find_assocmap, AssocMapExt.get_default. rewrite H6. auto. } - apply H6 in X2. apply H6 in X4. simplify. rewrite <- X2. rewrite <- X4. - apply int_eq_not. auto. - apply max_reg_stmnt_le_stmnt_tree in H2. - apply expr_lt_max_module_controllogic in H2. simplify. remember (max_reg_module m). lia. - apply max_reg_stmnt_le_stmnt_tree in H2. - apply expr_lt_max_module_controllogic in H2. simplify. remember (max_reg_module m). lia. - } - 2: { - simplify. unfold merge_regs. erewrite merge_get_default4. eauto. - repeat rewrite AssocMap.gso by lia. apply AssocMap.gss. - } - solve [auto]. - simplify. unfold merge_regs. erewrite merge_get_default4. eauto. - repeat rewrite AssocMap.gso by lia. apply AssocMap.gss. - simplify. unfold merge_regs. erewrite merge_get_default4. eauto. - repeat rewrite AssocMap.gso by lia. apply AssocMap.gss. - simplify. auto. - simplify. auto. - unfold merge_regs. repeat rewrite ! AssocMapExt.merge_add_assoc. - unfold empty_assocmap. - assert (mod_st (transf_module m) < max_reg_module m + 1). - { unfold max_reg_module, transf_module; repeat destruct_match; try discriminate; simplify; lia. } - remember (max_reg_module m). - repeat rewrite AssocMap.gso by lia. - unfold_merge. - apply merge_get_default4. - repeat rewrite AssocMap.gso by lia. - unfold transf_module; repeat destruct_match; try discriminate; simplify. - eapply match_assocmaps_merge in H21. - 2: { apply H21. } - admit. admit. - - econstructor. unfold merge_regs. admit. - - apply merge_arr_empty. apply match_empty_size_merge. - apply match_empty_assocmap_set. - eapply match_arrs_size_stmnt_preserved in H4; mgen_crush_local. - eapply match_arrs_size_stmnt_preserved in H4; mgen_crush_local. - apply match_arrs_merge_set2; auto. - eapply match_arrs_size_stmnt_preserved in H4; mgen_crush_local. - eapply match_arrs_size_stmnt_preserved in H4; mgen_crush_local. - eapply match_arrs_size_stmnt_preserved in H12; mgen_crush_local. - rewrite empty_stack_transf. mgen_crush_local. - eapply match_arrs_size_stmnt_preserved in H12; mgen_crush_local. - rewrite empty_stack_transf. mgen_crush_local. - auto. - admit. - admit. admit. - - unfold alt_load in *; simplify. inv H6. - 2: { match goal with H: context[location_is] |- _ => inv H end. } - match goal with H: context[location_is] |- _ => inv H end. - inv H30. simplify. inv H4. - 2: { match goal with H: context[location_is] |- _ => inv H end. } - inv H27. simplify. - do 2 econstructor. eapply Smallstep.plus_two. - econstructor. mgen_crush_local. mgen_crush_local. mgen_crush_local. eassumption. - eassumption. econstructor. simplify. econstructor. econstructor. - solve [eauto with mgen]. econstructor. econstructor. econstructor. - econstructor. econstructor. auto. auto. auto. - econstructor. econstructor. econstructor. - econstructor. econstructor. econstructor. cbn. eapply expr_runp_matches2; auto. eassumption. - 2: { eassumption. } - pose proof H3 as X. apply max_reg_stmnt_le_stmnt_tree in X. - apply expr_lt_max_module_datapath in X. simplify. remember (max_reg_module m); lia. - auto. - - simplify. rewrite empty_stack_transf. unfold transf_module; repeat destruct_match; crush. - eapply exec_ram_Some_read; simplify. - 2: { - unfold merge_regs. admit. - } - 2: { - unfold merge_regs. admit. - } - { unfold disable_ram, transf_module in DISABLE_RAM; - repeat destruct_match; try discriminate. simplify. apply int_eq_not. auto. admit. } (* } *) -(* { unfold merge_regs; unfold_merge. repeat rewrite AssocMap.gso by lia. apply AssocMap.gss. } *) -(* { unfold merge_regs; unfold_merge. apply AssocMap.gss. } *) -(* { eapply match_arrs_read. eassumption. mgen_crush_local. } *) -(* { crush. } *) -(* { crush. } *) -(* { unfold merge_regs. unfold_merge. *) -(* unfold transf_module; repeat destruct_match; try discriminate; simplify. *) -(* assert (mod_st m < max_reg_module m + 1). *) -(* { unfold max_reg_module; lia. } *) -(* remember (max_reg_module m). repeat rewrite AssocMap.gso by lia. *) -(* apply AssocMap.gss. *) -(* } *) -(* { auto. } *) - -(* { econstructor. *) -(* { unfold merge_regs. unfold_merge. *) -(* assert (mod_reset m < max_reg_module m + 1). *) -(* { unfold max_reg_module; lia. } *) -(* unfold transf_module; repeat destruct_match; try discriminate; simplify. *) -(* assert (mod_st m < mod_reset m). *) -(* { pose proof (mod_ordering_wf m); unfold module_ordering in *. simplify. *) -(* repeat match goal with *) -(* | H: context[_ apply Pos.ltb_lt in H *) -(* end; lia. *) -(* } *) -(* repeat rewrite AssocMap.gso by lia. *) -(* inv ASSOC. rewrite <- H19. auto. lia. *) -(* } *) -(* { unfold merge_regs. unfold_merge. *) -(* assert (mod_finish m < max_reg_module m + 1). *) -(* { unfold max_reg_module; lia. } *) -(* unfold transf_module; repeat destruct_match; try discriminate; simplify. *) -(* assert (mod_st m < mod_finish m). *) -(* { pose proof (mod_ordering_wf m). unfold module_ordering in *. simplify. *) -(* repeat match goal with *) -(* | H: context[_ apply Pos.ltb_lt in H *) -(* end; lia. *) -(* } *) -(* repeat rewrite AssocMap.gso by lia. *) -(* inv ASSOC. rewrite <- H19. auto. lia. *) -(* } *) -(* { unfold merge_regs. unfold_merge. *) -(* assert (mod_st m < max_reg_module m + 1). *) -(* { unfold max_reg_module; lia. } *) -(* unfold transf_module; repeat destruct_match; try discriminate; simplify. *) -(* repeat rewrite AssocMap.gso by lia. apply AssocMap.gss. *) -(* } *) -(* { eassumption. } *) -(* { eassumption. } *) -(* { econstructor. econstructor. simplify. unfold merge_regs. unfold_merge. *) -(* eapply expr_runp_matches. eassumption. *) -(* assert (max_reg_expr x0 + 1 <= max_reg_module m + 1). *) -(* { pose proof H2 as X. apply max_reg_stmnt_le_stmnt_tree in X. *) -(* apply expr_lt_max_module_controllogic in X. simplify. remember (max_reg_module m). lia. } *) -(* assert (max_reg_expr x0 + 1 <= mod_st m). *) -(* { unfold module_ordering in *. simplify. *) -(* repeat match goal with *) -(* | H: context[_ apply Pos.ltb_lt in H *) -(* end. *) -(* pose proof H2 as X. apply max_reg_stmnt_le_stmnt_tree in X. *) -(* simplify. lia. *) -(* } *) -(* remember (max_reg_module m). *) -(* apply match_assocmaps_gt; [lia|]. *) -(* apply match_assocmaps_gt; [lia|]. *) -(* apply match_assocmaps_gt; [lia|]. *) -(* apply match_assocmaps_gt; [lia|]. *) -(* apply match_assocmaps_gt; [lia|]. *) -(* apply match_assocmaps_gt; [lia|]. *) -(* simplify. *) -(* eapply match_assocmaps_ge. eauto. lia. *) -(* mgen_crush_local. *) -(* } *) -(* { simplify. unfold merge_regs. unfold_merge. *) -(* unfold transf_module; repeat destruct_match; try discriminate; simplify. *) -(* assert (mod_st m < max_reg_module m + 1). *) -(* { unfold max_reg_module; lia. } *) -(* remember (max_reg_module m). *) -(* repeat rewrite AssocMap.gso by lia. apply AssocMap.gss. *) -(* } *) -(* { *) -(* simplify. econstructor. econstructor. econstructor. simplify. *) -(* unfold merge_regs; unfold_merge. *) -(* repeat rewrite find_assocmap_gso by lia. apply find_assocmap_gss. *) -(* } *) -(* { simplify. rewrite empty_stack_transf. *) -(* unfold transf_module; repeat destruct_match; try discriminate; simplify. *) -(* econstructor. simplify. *) -(* unfold merge_regs; unfold_merge. simplify. *) -(* assert (r < max_reg_module m + 1). *) -(* { pose proof H3 as X. eapply max_reg_module_datapath_gt with (p := max_reg_module m + 1) in X. *) -(* unfold max_reg_stmnt in X. simplify. *) -(* lia. lia. } *) -(* assert (mod_st m < max_reg_module m + 1). *) -(* { unfold max_reg_module; lia. } *) -(* repeat rewrite find_assocmap_gso by lia. rewrite find_assocmap_gss. *) -(* repeat rewrite find_assocmap_gso by lia. rewrite find_assocmap_gss. *) -(* apply Int.eq_true. *) -(* } *) -(* { crush. } *) -(* { crush. } *) -(* { unfold merge_regs. unfold_merge. simplify. *) -(* assert (r < mod_st m). *) -(* { unfold module_ordering in *. simplify. *) -(* repeat match goal with *) -(* | H: context[_ apply Pos.ltb_lt in H *) -(* end. *) -(* pose proof H3 as X. apply max_reg_stmnt_le_stmnt_tree in X. *) -(* simplify. lia. *) -(* } *) -(* unfold merge_regs in H8. repeat rewrite AssocMapExt.merge_add_assoc in H8. *) -(* simplify. rewrite AssocMap.gso in H8 by lia. rewrite AssocMap.gss in H8. *) -(* unfold transf_module; repeat destruct_match; try discriminate; simplify. *) -(* repeat rewrite AssocMap.gso by lia. *) -(* apply AssocMap.gss. } *) -(* { eassumption. } *) -(* } *) -(* { eauto. } *) -(* { econstructor. *) -(* { unfold merge_regs. unfold_merge. simplify. *) -(* apply match_assocmaps_gss. *) -(* unfold merge_regs in H8. repeat rewrite AssocMapExt.merge_add_assoc in H8. *) -(* rewrite AssocMap.gso in H8. rewrite AssocMap.gss in H8. inv H8. *) -(* remember (max_reg_module m). *) -(* assert (mod_st m < max_reg_module m + 1). *) -(* { unfold max_reg_module; lia. } *) -(* apply match_assocmaps_switch_neq; [|lia]. *) -(* apply match_assocmaps_gt; [lia|]. *) -(* apply match_assocmaps_switch_neq; [|lia]. *) -(* apply match_assocmaps_gt; [lia|]. *) -(* apply match_assocmaps_switch_neq; [|lia]. *) -(* apply match_assocmaps_gt; [lia|]. *) -(* apply match_assocmaps_switch_neq; [|lia]. *) -(* apply match_assocmaps_gt; [lia|]. *) -(* apply match_assocmaps_switch_neq; [|lia]. *) -(* apply match_assocmaps_gt; [lia|]. *) -(* apply match_assocmaps_duplicate. *) -(* apply match_assocmaps_gss. auto. *) -(* assert (r < mod_st m). *) -(* { unfold module_ordering in *. simplify. *) -(* repeat match goal with *) -(* | H: context[_ apply Pos.ltb_lt in H *) -(* end. *) -(* pose proof H3 as X. apply max_reg_stmnt_le_stmnt_tree in X. *) -(* simplify. lia. *) -(* } lia. *) -(* } *) -(* { *) -(* apply merge_arr_empty. mgen_crush_local. *) -(* apply merge_arr_empty2. mgen_crush_local. *) -(* apply merge_arr_empty2. mgen_crush_local. *) -(* apply merge_arr_empty2. mgen_crush_local. *) -(* mgen_crush_local. *) -(* } *) -(* { auto. } *) -(* { mgen_crush_local. } *) -(* { mgen_crush_local. } *) -(* { unfold disable_ram. *) -(* unfold transf_module; repeat destruct_match; try discriminate; simplify. *) -(* unfold merge_regs. unfold_merge. simplify. *) -(* assert (mod_st m < max_reg_module m + 1). *) -(* { unfold max_reg_module; lia. } *) -(* assert (r < max_reg_module m + 1). *) -(* { pose proof H3 as X. eapply max_reg_module_datapath_gt with (p := max_reg_module m + 1) in X. *) -(* unfold max_reg_stmnt in X. simplify. *) -(* lia. lia. } *) -(* repeat rewrite find_assocmap_gso by lia. *) -(* rewrite find_assocmap_gss. *) -(* repeat rewrite find_assocmap_gso by lia. *) -(* rewrite find_assocmap_gss. apply Int.eq_true. *) -(* } *) -(* } *) -(* Qed. *) Admitted. - -Lemma exec_ram_resets_en : - forall rs ar rs' ar' r, - exec_ram rs ar (Some r) rs' ar' -> - assoc_nonblocking rs = empty_assocmap -> - Int.eq ((assoc_blocking (merge_reg_assocs rs')) # (ram_en r, 32)) - ((assoc_blocking (merge_reg_assocs rs')) # (ram_u_en r, 32)) = true. -Proof. - inversion 1; intros; subst; unfold merge_reg_assocs; simplify. - - rewrite H6. mgen_crush_local. - (* - unfold merge_regs. rewrite H12. unfold_merge. *) -(* unfold find_assocmap, AssocMapExt.get_default in *. *) -(* rewrite AssocMap.gss; auto. rewrite AssocMap.gso; auto. setoid_rewrite H4. apply Int.eq_true. *) -(* pose proof (ram_ordering r); lia. *) -(* - unfold merge_regs. rewrite H11. unfold_merge. *) -(* unfold find_assocmap, AssocMapExt.get_default in *. *) -(* rewrite AssocMap.gss; auto. *) -(* repeat rewrite AssocMap.gso by (pose proof (ram_ordering r); lia). *) -(* setoid_rewrite H3. apply Int.eq_true. *) -(* Qed. *) Admitted. - -Lemma disable_ram_set_gso : - forall rs r i v, - disable_ram (Some r) rs -> - i <> (ram_en r) -> i <> (ram_u_en r) -> - disable_ram (Some r) (rs # i <- v). -Proof. - unfold disable_ram, find_assocmap, AssocMapExt.get_default; intros; - repeat rewrite AssocMap.gso by lia; auto. -Qed. -#[local] Hint Resolve disable_ram_set_gso : mgen. - -Lemma disable_ram_None rs : disable_ram None rs. -Proof. unfold disable_ram; auto. Qed. -#[local] Hint Resolve disable_ram_None : mgen. - -Lemma init_regs_equal_empty l st : - Forall (Pos.gt st) l -> (init_regs nil l) ! st = None. -Proof. induction l; simplify; apply AssocMap.gempty. Qed. - -Lemma forall_lt_num : - forall l p p', Forall (Pos.gt p) l -> p < p' -> Forall (Pos.gt p') l. -Proof. induction l; crush; inv H; constructor; [lia | eauto]. Qed. - -Section CORRECTNESS. - - Context (prog tprog: program). - Context (TRANSL: match_prog prog tprog). - - Let ge : genv := Genv.globalenv prog. - Let tge : genv := Genv.globalenv tprog. - - Lemma symbols_preserved: - forall (s: AST.ident), Genv.find_symbol tge s = Genv.find_symbol ge s. - Proof using TRANSL. intros. eapply (Genv.find_symbol_match TRANSL). Qed. - #[local] Hint Resolve symbols_preserved : mgen. - - Lemma function_ptr_translated: - forall (b: Values.block) (f: fundef), - Genv.find_funct_ptr ge b = Some f -> - exists tf, - Genv.find_funct_ptr tge b = Some tf /\ transf_fundef f = tf. - Proof using TRANSL. - intros. exploit (Genv.find_funct_ptr_match TRANSL); eauto. - intros (cu & tf & P & Q & R); exists tf; auto. - Qed. - #[local] Hint Resolve function_ptr_translated : mgen. - - Lemma functions_translated: - forall (v: Values.val) (f: fundef), - Genv.find_funct ge v = Some f -> - exists tf, - Genv.find_funct tge v = Some tf /\ transf_fundef f = tf. - Proof using TRANSL. - intros. exploit (Genv.find_funct_match TRANSL); eauto. - intros (cu & tf & P & Q & R); exists tf; auto. - Qed. - #[local] Hint Resolve functions_translated : mgen. - - Lemma senv_preserved: - Senv.equiv (Genv.to_senv ge) (Genv.to_senv tge). - Proof (Genv.senv_transf TRANSL). - #[local] Hint Resolve senv_preserved : mgen. - - Theorem transf_step_correct: - forall (S1 : state) t S2, - step ge S1 t S2 -> - forall R1, - match_states S1 R1 -> - exists R2, Smallstep.plus step tge R1 t R2 /\ match_states S2 R2. - Proof. - Ltac transf_step_correct_assum := - match goal with - | H: match_states _ _ |- _ => let H2 := fresh "MATCH" in learn H as H2; inv H2 - | H: mod_ram ?m = Some ?r |- _ => - let H2 := fresh "RAM" in learn H; - pose proof (transf_module_code_ram m r H) as H2 - | H: mod_ram ?m = None |- _ => - let H2 := fresh "RAM" in learn H; - pose proof (transf_module_code m H) as H2 - end. - Ltac transf_step_correct_tac := - match goal with - | |- Smallstep.plus _ _ _ _ _ => apply Smallstep.plus_one - end. - induction 1; destruct (mod_ram m) eqn:RAM; simplify; repeat transf_step_correct_assum; - repeat transf_step_correct_tac. - - assert (MATCH_SIZE1: match_empty_size m nasa1 /\ match_empty_size m basa1). - { eapply match_arrs_size_stmnt_preserved2; eauto. unfold match_empty_size; mgen_crush_local. } - simplify. - assert (MATCH_SIZE2: match_empty_size m nasa2 /\ match_empty_size m basa2). - { eapply match_arrs_size_stmnt_preserved2; mgen_crush_local. } simplify. - assert (MATCH_SIZE2: match_empty_size m nasa3 /\ match_empty_size m basa3). - { eapply match_arrs_size_ram_preserved2; mgen_crush_local. } simplify. - assert (MATCH_ARR3: match_arrs_size nasa3 basa3) by mgen_crush_local. - exploit match_states_same. apply H4. eauto with mgen. - econstructor; eauto. econstructor; eauto. econstructor; eauto. econstructor; eauto. - intros. repeat inv_exists. simplify. inv H18. inv H21. - exploit match_states_same. apply H6. eauto with mgen. - econstructor; eauto. econstructor; eauto. intros. repeat inv_exists. simplify. inv H18. inv H23. - exploit exec_ram_same; eauto. eauto with mgen. - econstructor. eapply match_assocmaps_merge; eauto. eauto with mgen. - econstructor. - apply match_arrs_merge; eauto with mgen. eauto with mgen. - intros. repeat inv_exists. simplify. inv H18. inv H28. - econstructor; simplify. apply Smallstep.plus_one. econstructor. - mgen_crush_local. mgen_crush_local. mgen_crush_local. rewrite RAM0; eassumption. rewrite RAM0; eassumption. - rewrite RAM0. eassumption. mgen_crush_local. eassumption. rewrite RAM0 in H21. rewrite RAM0. - rewrite RAM. eassumption. eauto. eauto. eauto with mgen. eauto. - econstructor. mgen_crush_local. apply match_arrs_merge; mgen_crush_local. eauto. - apply match_empty_size_merge; mgen_crush_local; mgen_crush_local. - assert (MATCH_SIZE1': match_empty_size m ran'0 /\ match_empty_size m rab'0). - { eapply match_arrs_size_stmnt_preserved2; eauto. unfold match_empty_size; mgen_crush_local. } - simplify. - assert (MATCH_SIZE2': match_empty_size m ran'2 /\ match_empty_size m rab'2). - { eapply match_arrs_size_stmnt_preserved2; mgen_crush_local. } simplify. - assert (MATCH_SIZE2': match_empty_size m ran'4 /\ match_empty_size m rab'4). - { eapply match_arrs_size_ram_preserved2; mgen_crush_local. - unfold match_empty_size, transf_module, empty_stack. - repeat destruct_match; crush. mgen_crush_local. } - apply match_empty_size_merge; mgen_crush_local; mgen_crush_local. - unfold disable_ram. - unfold transf_module; repeat destruct_match; crush. - apply exec_ram_resets_en in H21. unfold merge_reg_assocs in H21. - simplify. auto. auto. - - eapply translation_correct; eauto. - - do 2 econstructor. apply Smallstep.plus_one. - apply step_finish; mgen_crush_local. constructor; eauto. - - do 2 econstructor. apply Smallstep.plus_one. - apply step_finish; mgen_crush_local. econstructor; eauto. - - econstructor. econstructor. apply Smallstep.plus_one. econstructor. - replace (mod_entrypoint (transf_module m)) with (mod_entrypoint m) by (rewrite RAM0; auto). - replace (mod_reset (transf_module m)) with (mod_reset m) by (rewrite RAM0; auto). - replace (mod_finish (transf_module m)) with (mod_finish m) by (rewrite RAM0; auto). - replace (empty_stack (transf_module m)) with (empty_stack m) by (rewrite RAM0; auto). - replace (mod_params (transf_module m)) with (mod_params m) by (rewrite RAM0; auto). - replace (mod_st (transf_module m)) with (mod_st m) by (rewrite RAM0; auto). - repeat econstructor; mgen_crush_local. - unfold disable_ram. unfold transf_module; repeat destruct_match; crush. - pose proof (mod_ordering_wf m); unfold module_ordering in *. - pose proof (mod_params_wf m). - pose proof (mod_ram_wf m r Heqo0). - pose proof (ram_ordering r). - simplify. - repeat rewrite find_assocmap_gso by lia. - assert ((init_regs nil (mod_params m)) ! (ram_en r) = None). - { apply init_regs_equal_empty. eapply forall_lt_num. eassumption. lia. } - assert ((init_regs nil (mod_params m)) ! (ram_u_en r) = None). - { apply init_regs_equal_empty. eapply forall_lt_num. eassumption. lia. } - unfold find_assocmap, AssocMapExt.get_default. rewrite H7. rewrite H14. auto. - - econstructor. econstructor. apply Smallstep.plus_one. econstructor. - replace (mod_entrypoint (transf_module m)) with (mod_entrypoint m). - replace (mod_reset (transf_module m)) with (mod_reset m). - replace (mod_finish (transf_module m)) with (mod_finish m). - replace (empty_stack (transf_module m)) with (empty_stack m). - replace (mod_params (transf_module m)) with (mod_params m). - replace (mod_st (transf_module m)) with (mod_st m). - all: try solve [unfold transf_module; repeat destruct_match; mgen_crush_local]. - repeat econstructor; mgen_crush_local. - unfold disable_ram. unfold transf_module; repeat destruct_match; crush. - unfold max_reg_module. - repeat rewrite find_assocmap_gso by lia. - assert (max_reg_module m + 1 > max_list (mod_params m)). - { unfold max_reg_module. lia. } - apply max_list_correct in H0. - unfold find_assocmap, AssocMapExt.get_default. - rewrite init_regs_equal_empty. rewrite init_regs_equal_empty. auto. - eapply forall_lt_num. eassumption. unfold max_reg_module. lia. - eapply forall_lt_num. eassumption. unfold max_reg_module. lia. - - inv STACKS. destruct b1; subst. - econstructor. econstructor. apply Smallstep.plus_one. - econstructor. eauto. - clear Learn. inv H0. inv H3. inv STACKS. inv H3. constructor. - constructor. intros. - rewrite RAM0. - destruct (Pos.eq_dec r res); subst. - rewrite AssocMap.gss. - rewrite AssocMap.gss. auto. - rewrite AssocMap.gso; auto. - symmetry. rewrite AssocMap.gso; auto. - destruct (Pos.eq_dec (mod_st m) r); subst. - rewrite AssocMap.gss. - rewrite AssocMap.gss. auto. - rewrite AssocMap.gso; auto. - symmetry. rewrite AssocMap.gso; auto. inv MATCH_ASSOC0. apply H1. auto. - auto. auto. auto. auto. - rewrite RAM0. rewrite RAM. rewrite RAM0 in DISABLE_RAM. rewrite RAM in DISABLE_RAM. - apply disable_ram_set_gso. - apply disable_ram_set_gso. auto. - pose proof (mod_ordering_wf m); unfold module_ordering in *. - pose proof (ram_ordering r0). simplify. - pose proof (mod_ram_wf m r0 H). lia. - pose proof (mod_ordering_wf m); unfold module_ordering in *. - pose proof (ram_ordering r0). simplify. - pose proof (mod_ram_wf m r0 H). lia. - pose proof (mod_ordering_wf m); unfold module_ordering in *. - pose proof (ram_ordering r0). simplify. - pose proof (mod_ram_wf m r0 H). lia. - pose proof (mod_ordering_wf m); unfold module_ordering in *. - pose proof (ram_ordering r0). simplify. - pose proof (mod_ram_wf m r0 H). lia. - - inv STACKS. destruct b1; subst. - econstructor. econstructor. apply Smallstep.plus_one. - econstructor. eauto. - clear Learn. inv H0. inv H3. inv STACKS. constructor. - constructor. intros. - unfold transf_module. repeat destruct_match; crush. - destruct (Pos.eq_dec r res); subst. - rewrite AssocMap.gss. - rewrite AssocMap.gss. auto. - rewrite AssocMap.gso; auto. - symmetry. rewrite AssocMap.gso; auto. - destruct (Pos.eq_dec (mod_st m) r); subst. - rewrite AssocMap.gss. - rewrite AssocMap.gss. auto. - rewrite AssocMap.gso; auto. - symmetry. rewrite AssocMap.gso; auto. inv MATCH_ASSOC. apply H. auto. - auto. auto. auto. auto. - Opaque disable_ram. - unfold transf_module in *; repeat destruct_match; crush. - apply disable_ram_set_gso. - apply disable_ram_set_gso. - auto. - simplify. unfold max_reg_module. lia. - simplify. unfold max_reg_module. lia. - simplify. unfold max_reg_module. lia. - simplify. unfold max_reg_module. lia. - Qed. - #[local] Hint Resolve transf_step_correct : mgen. - - Lemma transf_initial_states : - forall s1 : state, - initial_state prog s1 -> - exists s2 : state, - initial_state tprog s2 /\ match_states s1 s2. - Proof using TRANSL. - simplify. inv H. - exploit function_ptr_translated. eauto. intros. - inv H. inv H3. - econstructor. econstructor. econstructor. - eapply (Genv.init_mem_match TRANSL); eauto. - setoid_rewrite (Linking.match_program_main TRANSL). - rewrite symbols_preserved. eauto. - eauto. - econstructor. - Qed. - #[local] Hint Resolve transf_initial_states : mgen. - - Lemma transf_final_states : - forall (s1 : state) - (s2 : state) - (r : Int.int), - match_states s1 s2 -> - final_state s1 r -> - final_state s2 r. - Proof using TRANSL. - intros. inv H0. inv H. inv STACKS. unfold valueToInt. constructor. auto. - Qed. - #[local] Hint Resolve transf_final_states : mgen. - - Theorem transf_program_correct: - Smallstep.forward_simulation (semantics prog) (semantics tprog). - Proof using TRANSL. - eapply Smallstep.forward_simulation_plus; mgen_crush_local. - apply senv_preserved. - Qed. - -End CORRECTNESS. diff --git a/src/hls/Predicate.v b/src/hls/Predicate.v index 3cd9b0f..0dafa9e 100644 --- a/src/hls/Predicate.v +++ b/src/hls/Predicate.v @@ -596,15 +596,6 @@ Proof. simpl in *|-. cbn. unfold sat_lit. cbn. crush. Qed. -Lemma stseytin_and_correct2 : - forall cur p1 p2 fm c, - stseytin_and cur p1 p2 = fm -> - sat_formula fm c -> - sat_lit cur c <-> sat_lit p1 c /\ sat_lit p2 c. -Proof. - intros. split. intros. inv H1. unfold stseytin_and in *. - inv H0; try contradiction. Abort. - Lemma stseytin_or_correct : forall cur p1 p2 fm c, stseytin_or cur p1 p2 = fm -> @@ -620,47 +611,6 @@ Proof. simpl in *|-. cbn. unfold sat_lit. cbn. crush. Qed. -Lemma stseytin_or_correct2 : - forall cur p1 p2 fm c, - stseytin_or cur p1 p2 = fm -> - sat_formula fm c -> - sat_lit cur c <-> sat_lit p1 c \/ sat_lit p2 c. -Proof. Abort. - -Lemma xtseytin_correct : - forall p next l n fm c, - xtseytin next p = (n, l, fm) -> - sat_predicate p c = true <-> sat_formula ((l::nil)::fm) c. -Proof. - induction p. - - intros. simplify. destruct p. - inv H. split. - intros. destruct b. split; crush. - apply negb_true_iff in H. - split; crush. - intros. inv H. inv H0; try contradiction. - inv H. simplify. rewrite <- H0. - destruct b. - rewrite -> H0; auto. - rewrite -> H0; auto. - - admit. - - admit. - - intros. split. intros. simpl in H0. - apply andb_prop in H0. inv H0. - cbn in H. - repeat destruct_match; try discriminate; []. inv H. eapply IHp1 in Heqp. -(* eapply IHp2 in Heqp1. apply Heqp1 in H2. - apply Heqp in H1. inv H1. inv H2. - assert - (sat_formula - (((true, n1) :: bar l0 :: bar l1 :: nil) - :: (bar (true, n1) :: l0 :: nil) - :: (bar (true, n1) :: l1 :: nil) :: nil) c). - eapply stseytin_and_correct. unfold stseytin_and. eauto. - unfold sat_lit. simpl. admit. - inv H; try contradiction. inv H1; try contradiction. eauto.*) -Abort. - Fixpoint max_predicate (p: pred_op) : positive := match p with | Plit (b, p) => p @@ -670,17 +620,6 @@ Fixpoint max_predicate (p: pred_op) : positive := | Por a b => Pos.max (max_predicate a) (max_predicate b) end. -Definition tseytin (p: pred_op) : - {fm: formula | forall a, - sat_predicate p a = true <-> sat_formula fm a}. - refine ( - (match xtseytin (max_predicate p + 1) p as X - return xtseytin (max_predicate p + 1) p = X -> - {fm: formula | forall a, sat_predicate p a = true <-> sat_formula fm a} - with (m, n, fm) => fun H => exist _ ((n::nil) :: fm) _ - end) (eq_refl (xtseytin (max_predicate p + 1) p))). - intros. Abort. - Definition tseytin_simple (p: pred_op) : formula := let m := (max_predicate p + 1)%positive in let '(m, n, fm) := xtseytin m p in diff --git a/src/hls/Value.v b/src/hls/Value.v deleted file mode 100644 index 0d3ea66..0000000 --- a/src/hls/Value.v +++ /dev/null @@ -1,552 +0,0 @@ -(*(* - * Vericert: Verified high-level synthesis. - * Copyright (C) 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 . - *) - -(* begin hide *) -From bbv Require Import Word. -From bbv Require HexNotation WordScope. -From Coq Require Import ZArith.ZArith FSets.FMapPositive Lia. -From compcert Require Import lib.Integers common.Values. -From vericert Require Import Vericertlib. -(* end hide *) - -(** * Value - -A [value] is a bitvector with a specific size. We are using the implementation -of the bitvector by mit-plv/bbv, because it has many theorems that we can reuse. -However, we need to wrap it with an [Inductive] so that we can specify and match -on the size of the [value]. This is necessary so that we can easily store -[value]s of different sizes in a list or in a map. - -Using the default [word], this would not be possible, as the size is part of the type. *) - -Record value : Type := - mkvalue { - vsize: nat; - vword: word vsize - }. - -(** ** Value conversions - -Various conversions to different number types such as [N], [Z], [positive] and -[int], where the last one is a theory of integers of powers of 2 in CompCert. *) - -Definition wordToValue : forall sz : nat, word sz -> value := mkvalue. - -Definition valueToWord : forall v : value, word (vsize v) := vword. - -Definition valueToNat (v :value) : nat := - wordToNat (vword v). - -Definition natToValue sz (n : nat) : value := - mkvalue sz (natToWord sz n). - -Definition valueToN (v : value) : N := - wordToN (vword v). - -Definition NToValue sz (n : N) : value := - mkvalue sz (NToWord sz n). - -Definition ZToValue (s : nat) (z : Z) : value := - mkvalue s (ZToWord s z). - -Definition valueToZ (v : value) : Z := - wordToZ (vword v). - -Definition uvalueToZ (v : value) : Z := - uwordToZ (vword v). - -Definition posToValue sz (p : positive) : value := - ZToValue sz (Zpos p). - -Definition posToValueAuto (p : positive) : value := - let size := Pos.to_nat (Pos.size p) in - ZToValue size (Zpos p). - -Definition valueToPos (v : value) : positive := - Z.to_pos (uvalueToZ v). - -Definition intToValue (i : Integers.int) : value := - ZToValue Int.wordsize (Int.unsigned i). - -Definition valueToInt (i : value) : Integers.int := - Int.repr (uvalueToZ i). - -Definition ptrToValue (i : Integers.ptrofs) : value := - ZToValue Ptrofs.wordsize (Ptrofs.unsigned i). - -Definition valueToPtr (i : value) : Integers.ptrofs := - Ptrofs.repr (uvalueToZ i). - -Definition valToValue (v : Values.val) : option value := - match v with - | Values.Vint i => Some (intToValue i) - | Values.Vptr b off => if Z.eqb (Z.modulo (uvalueToZ (ptrToValue off)) 4) 0%Z - then Some (ptrToValue off) - else None - | Values.Vundef => Some (ZToValue 32 0%Z) - | _ => None - end. - -(** Convert a [value] to a [bool], so that choices can be made based on the -result. This is also because comparison operators will give back [value] instead -of [bool], so if they are in a condition, they will have to be converted before -they can be used. *) - -Definition valueToBool (v : value) : bool := - negb (weqb (@wzero (vsize v)) (vword v)). - -Definition boolToValue (sz : nat) (b : bool) : value := - natToValue sz (if b then 1 else 0). - -(** ** Arithmetic operations *) - -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. auto. Qed. - -Definition value_eq_size: - forall v1 v2 : value, { vsize v1 = vsize v2 } + { True }. -Proof. - intros; destruct (Nat.eqb (vsize v1) (vsize v2)) eqn:?. - left; apply Nat.eqb_eq in Heqb; assumption. - right; trivial. -Defined. - -Definition map_any {A : Type} (v1 v2 : value) (f : word (vsize v1) -> word (vsize v1) -> A) - (EQ : vsize v1 = vsize v2) : A := - let w2 := unify_word (vsize v1) (vsize v2) (vword v2) EQ in - f (vword v1) w2. - -Definition map_any_opt {A : Type} (sz : nat) (v1 v2 : value) (f : word (vsize v1) -> word (vsize v1) -> A) - : option A := - match value_eq_size v1 v2 with - | left EQ => - Some (map_any v1 v2 f EQ) - | _ => None - end. - -Definition map_word (f : forall sz : nat, word sz -> word sz) (v : value) : value := - mkvalue (vsize v) (f (vsize v) (vword v)). - -Definition map_word2 (f : forall sz : nat, word sz -> word sz -> word sz) (v1 v2 : value) - (EQ : (vsize v1 = vsize v2)) : value := - let w2 := unify_word (vsize v1) (vsize v2) (vword v2) EQ in - mkvalue (vsize v1) (f (vsize v1) (vword v1) w2). - -Definition map_word2_opt (f : forall sz : nat, word sz -> word sz -> word sz) (v1 v2 : value) - : option value := - match value_eq_size v1 v2 with - | left EQ => Some (map_word2 f v1 v2 EQ) - | _ => None - end. - -Definition eq_to_opt (v1 v2 : value) (f : vsize v1 = vsize v2 -> value) - : option value := - match value_eq_size v1 v2 with - | left EQ => Some (f EQ) - | _ => None - end. - -Lemma eqvalue {sz : nat} (x y : word sz) : x = y <-> mkvalue sz x = mkvalue sz y. -Proof. - split; intros. - subst. reflexivity. inversion H. apply existT_wordToZ in H1. - apply wordToZ_inj. assumption. -Qed. - -Lemma eqvaluef {sz : nat} (x y : word sz) : x = y -> mkvalue sz x = mkvalue sz y. -Proof. apply eqvalue. Qed. - -Lemma nevalue {sz : nat} (x y : word sz) : x <> y <-> mkvalue sz x <> mkvalue sz y. -Proof. split; intros; intuition. apply H. apply eqvalue. assumption. - apply H. rewrite H0. trivial. -Qed. - -Lemma nevaluef {sz : nat} (x y : word sz) : x <> y -> mkvalue sz x <> mkvalue sz y. -Proof. apply nevalue. Qed. - -(*Definition rewrite_word_size (initsz finalsz : nat) (w : word initsz) - : option (word finalsz) := - match Nat.eqb initsz finalsz return option (word finalsz) with - | true => Some _ - | false => None - end.*) - -Definition valueeq (sz : nat) (x y : word sz) : - {mkvalue sz x = mkvalue sz y} + {mkvalue sz x <> mkvalue sz y} := - match weq x y with - | left eq => left (eqvaluef x y eq) - | right ne => right (nevaluef x y ne) - end. - -Definition valueeqb (x y : value) : bool := - match value_eq_size x y with - | left EQ => - weqb (vword x) (unify_word (vsize x) (vsize y) (vword y) EQ) - | right _ => false - end. - -Definition value_projZ_eqb (v1 v2 : value) : bool := Z.eqb (valueToZ v1) (valueToZ v2). - -Theorem value_projZ_eqb_true : - forall v1 v2, - v1 = v2 -> value_projZ_eqb v1 v2 = true. -Proof. intros. subst. unfold value_projZ_eqb. apply Z.eqb_eq. trivial. Qed. - -Theorem valueeqb_true_iff : - forall v1 v2, - valueeqb v1 v2 = true <-> v1 = v2. -Proof. - split; intros. - unfold valueeqb in H. destruct (value_eq_size v1 v2) eqn:?. - - destruct v1, v2. simpl in H. -Abort. - -Definition value_int_eqb (v : value) (i : int) : bool := - Z.eqb (valueToZ v) (Int.unsigned i). - -(** Arithmetic operations over [value], interpreting them as signed or unsigned -depending on the operation. - -The arithmetic operations over [word] are over [N] by default, however, can also -be called over [Z] explicitly, which is where the bits are interpreted in a -signed manner. *) - -Definition vplus v1 v2 := map_word2 wplus v1 v2. -Definition vplus_opt v1 v2 := map_word2_opt wplus v1 v2. -Definition vminus v1 v2 := map_word2 wminus v1 v2. -Definition vmul v1 v2 := map_word2 wmult v1 v2. -Definition vdiv v1 v2 := map_word2 wdiv v1 v2. -Definition vmod v1 v2 := map_word2 wmod v1 v2. - -Definition vmuls v1 v2 := map_word2 wmultZ v1 v2. -Definition vdivs v1 v2 := map_word2 wdivZ v1 v2. -Definition vmods v1 v2 := map_word2 wremZ v1 v2. - -(** ** Bitwise operations - -Bitwise operations over [value], which is independent of whether the number is -signed or unsigned. *) - -Definition vnot v := map_word wnot v. -Definition vneg v := map_word wneg v. -Definition vbitneg v := boolToValue (vsize v) (negb (valueToBool v)). -Definition vor v1 v2 := map_word2 wor v1 v2. -Definition vand v1 v2 := map_word2 wand v1 v2. -Definition vxor v1 v2 := map_word2 wxor v1 v2. - -(** ** Comparison operators - -Comparison operators that return a bool, there should probably be an equivalent -which returns another number, however I might just add that as an explicit -conversion. *) - -Definition veqb v1 v2 := map_any v1 v2 (@weqb (vsize v1)). -Definition vneb v1 v2 EQ := negb (veqb v1 v2 EQ). - -Definition veq v1 v2 EQ := boolToValue (vsize v1) (veqb v1 v2 EQ). -Definition vne v1 v2 EQ := boolToValue (vsize v1) (vneb v1 v2 EQ). - -Definition vltb v1 v2 := map_any v1 v2 wltb. -Definition vleb v1 v2 EQ := negb (map_any v2 v1 wltb (eq_sym EQ)). -Definition vgtb v1 v2 EQ := map_any v2 v1 wltb (eq_sym EQ). -Definition vgeb v1 v2 EQ := negb (map_any v1 v2 wltb EQ). - -Definition vltsb v1 v2 := map_any v1 v2 wsltb. -Definition vlesb v1 v2 EQ := negb (map_any v2 v1 wsltb (eq_sym EQ)). -Definition vgtsb v1 v2 EQ := map_any v2 v1 wsltb (eq_sym EQ). -Definition vgesb v1 v2 EQ := negb (map_any v1 v2 wsltb EQ). - -Definition vlt v1 v2 EQ := boolToValue (vsize v1) (vltb v1 v2 EQ). -Definition vle v1 v2 EQ := boolToValue (vsize v1) (vleb v1 v2 EQ). -Definition vgt v1 v2 EQ := boolToValue (vsize v1) (vgtb v1 v2 EQ). -Definition vge v1 v2 EQ := boolToValue (vsize v1) (vgeb v1 v2 EQ). - -Definition vlts v1 v2 EQ := boolToValue (vsize v1) (vltsb v1 v2 EQ). -Definition vles v1 v2 EQ := boolToValue (vsize v1) (vlesb v1 v2 EQ). -Definition vgts v1 v2 EQ := boolToValue (vsize v1) (vgtsb v1 v2 EQ). -Definition vges v1 v2 EQ := boolToValue (vsize v1) (vgesb v1 v2 EQ). - -(** ** Shift operators - -Shift operators on values. *) - -Definition shift_map (sz : nat) (f : word sz -> nat -> word sz) (w1 w2 : word sz) := - f w1 (wordToNat w2). - -Definition vshl v1 v2 := map_word2 (fun sz => shift_map sz (@wlshift sz)) v1 v2. -Definition vshr v1 v2 := map_word2 (fun sz => shift_map sz (@wrshift sz)) v1 v2. - -Module HexNotationValue. - Export HexNotation. - Import WordScope. - - Notation "sz ''h' a" := (NToValue sz (hex a)) (at level 50). - -End HexNotationValue. - -Inductive val_value_lessdef: val -> value -> Prop := -| val_value_lessdef_int: - forall i v', - i = valueToInt v' -> - val_value_lessdef (Vint i) v' -| val_value_lessdef_ptr: - forall b off v', - off = valueToPtr v' -> - (Z.modulo (uvalueToZ 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 := -| 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 -> - valueToZ (ZToValue (S n) z) = z. -Proof. - unfold valueToZ, ZToValue. simpl. - auto using wordToZ_ZToWord. -Qed. - -Lemma uvalueToZ_ZToValue : - forall n z, - (0 <= z < 2 ^ Z.of_nat n)%Z -> - uvalueToZ (ZToValue n z) = z. -Proof. - unfold uvalueToZ, ZToValue. simpl. - 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. -Proof. - intros. - unfold ZToValue, uvalueToZ. - rewrite ZToWord_uwordToZ. destruct v; auto. -Qed. - -Lemma valueToPos_posToValueAuto : - 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 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. -Proof. - intros. - unfold valueToInt, intToValue. rewrite uvalueToZ_ZToValue. auto using Int.repr_unsigned. - split. apply Int.unsigned_range_2. - assert ((Int.unsigned v <= Int.max_unsigned)%Z) by apply Int.unsigned_range_2. - apply Z.lt_le_pred in H. apply H. -Qed. - -Lemma valueToPtr_ptrToValue : - forall v, - valueToPtr (ptrToValue v) = v. -Proof. - intros. - unfold valueToPtr, ptrToValue. rewrite uvalueToZ_ZToValue. auto using Ptrofs.repr_unsigned. - split. apply Ptrofs.unsigned_range_2. - assert ((Ptrofs.unsigned v <= Ptrofs.max_unsigned)%Z) by apply Ptrofs.unsigned_range_2. - apply Z.lt_le_pred in H. apply H. -Qed. - -Lemma intToValue_valueToInt : - forall v, - vsize v = 32%nat -> - intToValue (valueToInt v) = v. -Proof. - intros. unfold valueToInt, intToValue. rewrite Int.unsigned_repr_eq. - unfold ZToValue, uvalueToZ. unfold Int.modulus. unfold Int.wordsize. unfold Wordsize_32.wordsize. - pose proof (uwordToZ_bound (vword v)). - rewrite Z.mod_small. rewrite <- H. rewrite ZToWord_uwordToZ. destruct v; auto. - rewrite <- H. rewrite two_power_nat_equiv. apply H0. -Qed. - -Lemma ptrToValue_valueToPtr : - forall v, - vsize v = 32%nat -> - ptrToValue (valueToPtr v) = v. -Proof. - intros. unfold valueToPtr, ptrToValue. rewrite Ptrofs.unsigned_repr_eq. - unfold ZToValue, uvalueToZ. unfold Ptrofs.modulus. unfold Ptrofs.wordsize. unfold Wordsize_Ptrofs.wordsize. - pose proof (uwordToZ_bound (vword v)). - rewrite Z.mod_small. rewrite <- H. rewrite ZToWord_uwordToZ. destruct v; auto. - rewrite <- H. rewrite two_power_nat_equiv. apply H0. -Qed. - -Lemma valToValue_lessdef : - forall v v', - valToValue v = Some v' -> - val_value_lessdef v v'. -Proof. - intros. - destruct v; try discriminate; constructor. - unfold valToValue in H. inversion H. - symmetry. apply valueToInt_intToValue. - inv H. destruct (uvalueToZ (ptrToValue i) mod 4 =? 0); try discriminate. - inv H1. symmetry. apply valueToPtr_ptrToValue. - inv H. destruct (uvalueToZ (ptrToValue i) mod 4 =? 0) eqn:?; try discriminate. - inv H1. apply Z.eqb_eq. apply Heqb0. -Qed. - -Lemma boolToValue_ValueToBool : - forall b, - valueToBool (boolToValue 32 b) = b. -Proof. destruct b; auto. Qed. - -Local Open Scope Z. - -Ltac word_op_value H := - intros; unfold uvalueToZ, ZToValue; simpl; rewrite unify_word_unfold; - rewrite <- H; rewrite uwordToZ_ZToWord_full; auto; omega. - -Lemma zadd_vplus : - 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, - vplus (ZToValue 32 z1) (ZToValue 32 z2) eq_refl = ZToValue 32 (z1 + z2). -Proof. - intros. unfold vplus, ZToValue, map_word2. rewrite unify_word_unfold. simpl. - rewrite ZToWord_plus; auto. -Qed. - -Lemma ZToValue_eq : - forall w1, - (mkvalue 32 w1) = (ZToValue 32 (wordToZ w1)). Abort. - -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_vplus. - rewrite <- Int.unsigned_repr_eq. - rewrite Int.repr_unsigned. auto. rewrite wordsize_32. omega. -Qed. - -(*Lemma intadd_vplus2 : - forall v1 v2 EQ, - vsize v1 = 32%nat -> - Int.add (valueToInt v1) (valueToInt v2) = valueToInt (vplus v1 v2 EQ). -Proof. - intros. unfold Int.add, valueToInt, intToValue. repeat (rewrite Int.unsigned_repr). - rewrite (@vadd_vplus v1 v2 EQ). trivial. - unfold uvalueToZ. pose proof (@uwordToZ_bound (vsize v2) (vword v2)). - rewrite H in EQ. rewrite <- EQ in H0 at 3.*) - (*rewrite zadd_vplus3. trivia*) - -Lemma valadd_vplus : - forall v1 v2 v1' v2' v v' EQ, - val_value_lessdef v1 v1' -> - val_value_lessdef v2 v2' -> - Val.add v1 v2 = v -> - vplus v1' v2' EQ = v' -> - val_value_lessdef v v'. -Proof. - intros. inv H; inv H0; constructor; simplify. - Abort. - -Lemma zsub_vminus : - 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 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 : - 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, - (sz > 0)%nat -> - (0 <= x < 2^(Z.of_nat sz))%Z -> - valueToNat (ZToValue sz x) = Z.to_nat x. -Proof. - destruct x; intros; unfold ZToValue, valueToNat; crush. - - rewrite wzero'_def. apply wordToNat_wzero. - - rewrite posToWord_nat. rewrite wordToNat_natToWord_2. trivial. - clear H1. - lazymatch goal with - | [ H : context[(_ < ?x)%Z] |- _ ] => replace x with (Z.of_nat (Z.to_nat x)) in H - end. - 2: { apply Z2Nat.id; apply Z.pow_nonneg; lia. } - - rewrite Z2Nat.inj_pow in H2; crush. - replace (Pos.to_nat 2) with 2%nat in H2 by reflexivity. - rewrite Nat2Z.id in H2. - rewrite <- positive_nat_Z in H2. - apply Nat2Z.inj_lt in H2. - assumption. -Qed. -*) -- cgit