From 5b5682a4d4ebb102002616015adab46154597b10 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Wed, 16 Dec 2020 15:52:48 +0100 Subject: upgrade kvx backend to coq.8.12.2 --- .gitignore | 1 + configure | 2 +- kvx/Asmblockdeps.v | 21 +++++++++++++-------- kvx/abstractbb/AbstractBasicBlocksDef.v | 15 +++++++++------ kvx/abstractbb/Parallelizability.v | 8 ++++---- kvx/lib/ForwardSimulationBlock.v | 26 +++++++++++++------------- kvx/lib/Machblockgen.v | 4 ++-- 7 files changed, 43 insertions(+), 34 deletions(-) diff --git a/.gitignore b/.gitignore index 771654d7..1d6cad94 100644 --- a/.gitignore +++ b/.gitignore @@ -21,6 +21,7 @@ # Emacs saves *~ # Executables and configuration +/tools/compiler_expand /ccomp /ccomp.byte /ccomp.prof diff --git a/configure b/configure index 415dbb03..e63ff928 100755 --- a/configure +++ b/configure @@ -565,7 +565,7 @@ missingtools=false echo "Testing Coq... " | tr -d '\n' coq_ver=$(${COQBIN}coqc -v 2>/dev/null | sed -n -e 's/The Coq Proof Assistant, version \([^ ]*\).*$/\1/p') case "$coq_ver" in - 8.8.0|8.8.1|8.8.2|8.9.0|8.9.1|8.10.0|8.10.1|8.10.2|8.11.0|8.11.1|8.11.2|8.12.0|8.12.1) + 8.8.0|8.8.1|8.8.2|8.9.0|8.9.1|8.10.0|8.10.1|8.10.2|8.11.0|8.11.1|8.11.2|8.12.0|8.12.1|8.12.2) echo "version $coq_ver -- good!";; ?*) echo "version $coq_ver -- UNSUPPORTED" diff --git a/kvx/Asmblockdeps.v b/kvx/Asmblockdeps.v index 3d981100..67430b54 100644 --- a/kvx/Asmblockdeps.v +++ b/kvx/Asmblockdeps.v @@ -40,6 +40,10 @@ Require Import Chunks. Require Import Lia. + +Import ListNotations. +Local Open Scope list_scope. + Open Scope impure. (** Definition of [L] *) @@ -706,19 +710,20 @@ Proof. destruct r; destruct r'. all: try discriminate; try contradiction. - intros. apply not_eq_add. apply ireg_to_pos_discr. congruence. - - intros. unfold ppos. cutrewrite (3 + ireg_to_pos g = (1 + ireg_to_pos g) + 2). apply Pos.add_no_neutral. - apply eq_sym. rewrite Pos.add_comm. rewrite Pos.add_assoc. reflexivity. + - intros. unfold ppos. replace (3 + ireg_to_pos g) with ((1 + ireg_to_pos g) + 2). + apply Pos.add_no_neutral. + rewrite Pos.add_comm, Pos.add_assoc. reflexivity. - intros. unfold ppos. rewrite Pos.add_comm. apply Pos.add_no_neutral. - intros. unfold ppos. apply not_eq_sym. - cutrewrite (3 + ireg_to_pos g = (1 + ireg_to_pos g) + 2). apply Pos.add_no_neutral. - apply eq_sym. rewrite Pos.add_comm. rewrite Pos.add_assoc. reflexivity. + replace (3 + ireg_to_pos g) with ((1 + ireg_to_pos g) + 2). apply Pos.add_no_neutral. + rewrite Pos.add_comm, Pos.add_assoc. reflexivity. - intros. unfold ppos. apply not_eq_sym. rewrite Pos.add_comm. apply Pos.add_no_neutral. Qed. Lemma ppos_pmem_discr: forall r, pmem <> ppos r. Proof. intros. destruct r. - - unfold ppos. unfold pmem. apply not_eq_sym. rewrite Pos.add_comm. cutrewrite (3 = 2 + 1). rewrite Pos.add_assoc. apply Pos.add_no_neutral. + - unfold ppos. unfold pmem. apply not_eq_sym. rewrite Pos.add_comm. replace 3 with (2 + 1). rewrite Pos.add_assoc. apply Pos.add_no_neutral. reflexivity. - unfold ppos. unfold pmem. discriminate. - unfold ppos. unfold pmem. discriminate. @@ -1250,7 +1255,7 @@ Theorem bisimu_par_exit ex sz ge fn rsr rsw mr mw sr sw: Proof. intros; unfold estep. exploit (bisimu_par_control ex sz rsw#PC ge fn rsr rsw mr mw sr sw); eauto. - cutrewrite (rsw # PC <- (rsw PC) = rsw); auto. + replace (rsw # PC <- (rsw PC)) with rsw; auto. apply extensionality. intros; destruct x; simpl; auto. Qed. @@ -1461,8 +1466,8 @@ Proof. destruct H2 as (m2' & H2 & H4). discriminate. rewrite H2 in H3. destruct (exec_bblock ge fn p2 rs m); simpl in H3. * destruct H3 as (s' & H3 & H5 & H6). inv H3. inv MS'. - cutrewrite (rs0=rs1). - - cutrewrite (m0=m1); auto. congruence. + replace rs0 with rs1. + - replace m0 with m1; auto. congruence. - apply functional_extensionality. intros r. generalize (H0 r). intros Hr. congruence. * discriminate. diff --git a/kvx/abstractbb/AbstractBasicBlocksDef.v b/kvx/abstractbb/AbstractBasicBlocksDef.v index 6960f363..34d72de1 100644 --- a/kvx/abstractbb/AbstractBasicBlocksDef.v +++ b/kvx/abstractbb/AbstractBasicBlocksDef.v @@ -273,23 +273,26 @@ with list_term := Scheme term_mut := Induction for term Sort Prop with list_term_mut := Induction for list_term Sort Prop. +Declare Scope pattern_scope. +Declare Scope term_scope. Bind Scope pattern_scope with term. Delimit Scope term_scope with term. Delimit Scope pattern_scope with pattern. +Local Open Scope pattern_scope. + Notation "[ ]" := (LTnil _) (format "[ ]"): pattern_scope. -Notation "[ x ]" := (LTcons x [] _): pattern_scope. +Notation "[ x ]" := (LTcons x [ ] _): pattern_scope. Notation "[ x ; y ; .. ; z ]" := (LTcons x (LTcons y .. (LTcons z (LTnil _) _) .. _) _): pattern_scope. Notation "o @ l" := (App o l _) (at level 50, no associativity): pattern_scope. Import HConsingDefs. Notation "[ ]" := (LTnil unknown_hid) (format "[ ]"): term_scope. -Notation "[ x ]" := (LTcons x [] unknown_hid): term_scope. +Notation "[ x ]" := (LTcons x []%term unknown_hid): term_scope. Notation "[ x ; y ; .. ; z ]" := (LTcons x (LTcons y .. (LTcons z (LTnil unknown_hid) unknown_hid) .. unknown_hid) unknown_hid): term_scope. Notation "o @ l" := (App o l unknown_hid) (at level 50, no associativity): term_scope. -Local Open Scope pattern_scope. Fixpoint term_eval (ge: genv) (t: term) (m: mem): option value := match t with @@ -370,7 +373,7 @@ Qed. Hint Resolve intro_fail_correct: wlp. (** The default reduction of a term to a pseudo-term *) -Definition identity_fail (t: term):= intro_fail [t] t. +Definition identity_fail (t: term):= intro_fail (t::nil) t. Lemma identity_fail_correct (t: term): match_pt t (identity_fail t). Proof. @@ -382,8 +385,8 @@ Hint Resolve identity_fail_correct: wlp. (** The reduction for constant term *) Definition nofail (is_constant: op -> bool) (t: term):= match t with - | Input x _ => intro_fail ([])%list t - | o @ [] => if is_constant o then (intro_fail ([])%list t) else (identity_fail t) + | Input x _ => intro_fail nil t + | o @ [] => if is_constant o then (intro_fail nil t) else (identity_fail t) | _ => identity_fail t end. diff --git a/kvx/abstractbb/Parallelizability.v b/kvx/abstractbb/Parallelizability.v index e5d21434..afa4b9fd 100644 --- a/kvx/abstractbb/Parallelizability.v +++ b/kvx/abstractbb/Parallelizability.v @@ -299,7 +299,7 @@ Proof. induction i as [|[y e] i']; cbn. - intros m tmp H x H0; inversion_clear H; auto. - intros m tmp H x (H1 & H2); destruct (exp_eval ge e tmp old); cbn; try congruence. - cutrewrite (m x = assign m y v x); eauto. + replace (m x) with (assign m y v x); eauto. rewrite assign_diff; auto. Qed. @@ -514,13 +514,13 @@ Proof. induction i as [|[x e] i']; cbn; auto. intros m tmp1 tmp2; rewrite disjoint_cons_l, disjoint_app_l. intros (H1 & H2 & H3) H6 H7. - cutrewrite (exp_eval ge e tmp1 old1 = exp_eval ge e tmp2 old2). + replace (exp_eval ge e tmp1 old1) with (exp_eval ge e tmp2 old2). - destruct (exp_eval ge e tmp2 old2); auto. eapply IHi'; eauto. cbn; intros x0 H0; unfold assign. destruct (R.eq_dec x x0); cbn; auto. - unfold disjoint in H2; apply exp_frame_correct. - intros;apply H6; auto. - intros;apply H7; auto. + intros;rewrite H6; auto. + intros;rewrite H7; auto. Qed. (** * Parallelizability *) diff --git a/kvx/lib/ForwardSimulationBlock.v b/kvx/lib/ForwardSimulationBlock.v index 61466dad..cc6ecdd8 100644 --- a/kvx/lib/ForwardSimulationBlock.v +++ b/kvx/lib/ForwardSimulationBlock.v @@ -25,7 +25,7 @@ Require Import Coqlib. Require Import Events. Require Import Globalenvs. Require Import Smallstep. - +Require Import Lia. Local Open Scope nat_scope. @@ -43,8 +43,8 @@ Lemma starN_split n s t s': exists (t1 t2:trace) s0, starN (step L) (globalenv L) m s t1 s0 /\ starN (step L) (globalenv L) k s0 t2 s' /\ t=t1**t2. Proof. induction 1; cbn. - + intros m k H; assert (X: m=0); try omega. - assert (X0: k=0); try omega. + + intros m k H; assert (X: m=0); try lia. + assert (X0: k=0); try lia. subst; repeat (eapply ex_intro); intuition eauto. + intros m; destruct m as [| m']; cbn. - intros k H2; subst; repeat (eapply ex_intro); intuition eauto. @@ -119,10 +119,10 @@ Proof. intros t [H1 H2] H3 H4. destruct (simu_mid_block _ _ _ H3 H4) as [H5 H6]; subst. constructor 1. - + omega. - + cutrewrite (dist_end_block head - dist_end_block next = S (dist_end_block head - dist_end_block previous)). + + lia. + + replace (dist_end_block head - dist_end_block next) with (S (dist_end_block head - dist_end_block previous)). - eapply starN_tailstep; eauto. - - omega. + - lia. Qed. Lemma follows_in_block_init (head current: state L1): @@ -131,10 +131,10 @@ Proof. intros t H3 H4. destruct (simu_mid_block _ _ _ H3 H4) as [H5 H6]; subst. constructor 1. - + omega. - + cutrewrite (dist_end_block head - dist_end_block current = 1). + + lia. + + replace (dist_end_block head - dist_end_block current) with 1. - eapply starN_tailstep; eauto. - - omega. + - lia. Qed. @@ -156,10 +156,10 @@ Proof. destruct s as [rs ms Hs]. cbn. destruct ms as [ms|]; unfold head; cbn; auto. constructor 1. - omega. - cutrewrite ((dist_end_block rs - dist_end_block rs)%nat=O). + lia. + replace (dist_end_block rs - dist_end_block rs)%nat with O. + apply starN_refl; auto. - + omega. + + lia. Qed. Inductive is_well_memorized (s s': memostate): Prop := @@ -259,7 +259,7 @@ Proof. constructor 1. destruct (simu_end_block (head s1) t (real s1') s2) as (s2' & H3 & H4); auto. * destruct (head_followed s1) as [H4 H3]. - cutrewrite (dist_end_block (head s1) - dist_end_block (real s1) = dist_end_block (head s1)) in H3; try omega. + replace (dist_end_block (head s1) - dist_end_block (real s1)) with (dist_end_block (head s1)) in H3; try lia. eapply starN_tailstep; eauto. * unfold head; rewrite H2; cbn. intuition eauto. Qed. diff --git a/kvx/lib/Machblockgen.v b/kvx/lib/Machblockgen.v index 3d5d7b2c..5a2f2a61 100644 --- a/kvx/lib/Machblockgen.v +++ b/kvx/lib/Machblockgen.v @@ -155,11 +155,11 @@ Proof. + intros; eapply Tr_add_label; eauto. destruct i; cbn in * |- *; congruence. + intros. remember (header bh0) as hbh0. destruct hbh0 as [|b]. * eapply Tr_add_basic; eauto. - * cutrewrite (add_basic bi empty_bblock = add_to_new_bblock (MB_basic bi)); auto. + * replace (add_basic bi empty_bblock) with (add_to_new_bblock (MB_basic bi)); auto. rewrite Heqti; eapply Tr_end_block; eauto. rewrite <- Heqti. eapply End_basic. congruence. + intros. - cutrewrite (cfi_bblock cfi = add_to_new_bblock (MB_cfi cfi)); auto. + replace (cfi_bblock cfi) with (add_to_new_bblock (MB_cfi cfi)); auto. rewrite Heqti. eapply Tr_end_block; eauto. rewrite <- Heqti. eapply End_cfi. congruence. Qed. -- cgit