diff --git a/.github/workflows/coq-action.yml b/.github/workflows/coq-action.yml index 7a72eee..4118182 100644 --- a/.github/workflows/coq-action.yml +++ b/.github/workflows/coq-action.yml @@ -1,6 +1,7 @@ name: CI on: + workflow_dispatch: push: branches: ['main'] pull_request: @@ -18,6 +19,7 @@ jobs: - '8.15' - '8.16' - '8.17' + - '8.18' - 'dev' ocaml_version: - 'default' @@ -29,4 +31,3 @@ jobs: opam_file: 'coq-quantumlib.opam' coq_version: ${{ matrix.coq_version }} ocaml_version: ${{ matrix.ocaml_version }} - diff --git a/DiscreteProb.v b/DiscreteProb.v index 30d905d..9aea7e0 100644 --- a/DiscreteProb.v +++ b/DiscreteProb.v @@ -102,7 +102,7 @@ Lemma sample_ub_lt : forall l r, Proof. induction l; intros. unfold sum_over_list in H. simpl in H. lra. simpl. destruct (Rlt_le_dec r a). lia. - apply lt_n_S. apply IHl. rewrite sum_over_list_cons in H. lra. + rewrite <- Nat.succ_lt_mono. apply IHl. rewrite sum_over_list_cons in H. lra. Qed. Lemma sample_lb : forall l r, (0 <= sample l r)%nat. @@ -1498,14 +1498,16 @@ Proof. replace (2 ^ (n + k))%nat with (2 ^ n * 2 ^ k)%nat by unify_pows_two. apply (Nat.mul_le_mono_r _ _ (2^k)%nat) in H0. rewrite Nat.mul_sub_distr_r in H0. - rewrite mult_1_l in H0. + rewrite Nat.mul_1_l in H0. apply (Nat.add_lt_le_mono x0 (2^k)%nat) in H0; auto. - rewrite <- le_plus_minus in H0. - rewrite plus_comm in H0. + rewrite (Nat.add_comm (2 ^ k) (2 ^ n * 2 ^ k - 2 ^ k)) in H0. + rewrite Nat.sub_add in H0. + rewrite Nat.add_comm in H0. assumption. - destruct (mult_O_le (2^k) (2^n))%nat; auto. + destruct (Arith_prebase.mult_O_le_stt (2^k) (2^n))%nat; auto. + destruct (Arith_prebase.mult_O_le_stt (2^k) (2^n))%nat; auto. assert (2 <> 0)%nat by lia. - apply (pow_positive 2 n) in H2. + apply (pow_positive 2 n) in H3. lia. - rewrite big_sum_0_bounded. rewrite Rmult_0_l; easy. diff --git a/Eigenvectors.v b/Eigenvectors.v index 82e34bd..a3d8577 100644 --- a/Eigenvectors.v +++ b/Eigenvectors.v @@ -6,7 +6,6 @@ Require Export Complex. Require Export Quantum. Require Import FTA. - (****************************) (** * Proving some indentities *) (****************************) @@ -40,7 +39,8 @@ Lemma XH_eq_HZ : σx × hadamard = hadamard × σz. Proof. lma'. Qed. Lemma SX_eq_YS : Sgate × σx = σy × Sgate. Proof. lma'; unfold Mmult; simpl; rewrite Cexp_PI2; lca. Qed. Lemma SZ_eq_ZS : Sgate × σz = σz × Sgate. Proof. lma'; unfold Mmult; - simpl; rewrite Cexp_PI2; lca. Qed. + simpl; rewrite Cexp_PI2; lca. Qed. + Lemma cnotX1 : cnot × (σx ⊗ I 2) = (σx ⊗ σx) × cnot. Proof. lma'. Qed. Lemma cnotX2 : cnot × (I 2 ⊗ σx) = (I 2 ⊗ σx) × cnot. Proof. lma'. Qed. Lemma cnotZ1 : cnot × (σz ⊗ I 2) = (σz ⊗ I 2) × cnot. Proof. lma'. Qed. @@ -258,8 +258,8 @@ Proof. induction n as [| n']. assert (H1 := (IHn' (reduce A 0 0))). rewrite p in H1; easy. } rewrite H', Pplus_comm, Pplus_degree2; auto. - rewrite H1. - apply le_lt_n_Sm. + rewrite H1. + rewrite Nat.lt_succ_r. apply Psum_degree; intros. assert (H2 : prep_mat A (S i) 0 = [A (S i) 0]). { unfold prep_mat. @@ -885,7 +885,7 @@ Proof. induction m2 as [| m2']. exists Zero. split. easy. rewrite smash_zero; try easy. - rewrite plus_0_r. + rewrite Nat.add_0_r. apply H2. - intros. rewrite (split_col T2) in *. @@ -1180,8 +1180,9 @@ Proof. intros a b m H0 Hdiv Hmod. rewrite Hdiv in Hmod. assert (H : m * (b / m) + (a - m * (b / m)) = m * (b / m) + (b - m * (b / m))). { rewrite Hmod. reflexivity. } - rewrite <- (le_plus_minus (m * (b / m)) a) in H. - rewrite <- (le_plus_minus (m * (b / m)) b) in H. + rewrite (Nat.add_comm (m * (b / m)) (a - m * (b / m))) in H. + rewrite (Nat.add_comm (m * (b / m)) (b - m * (b / m))) in H. + rewrite ! Nat.sub_add in H. apply H. apply Nat.mul_div_le; apply H0. rewrite <- Hdiv; apply Nat.mul_div_le; apply H0. diff --git a/Matrix.v b/Matrix.v index 9ffb90e..0c41802 100644 --- a/Matrix.v +++ b/Matrix.v @@ -306,8 +306,8 @@ Ltac by_cell := let Hi := fresh "Hi" in let Hj := fresh "Hj" in intros i j Hi Hj; try solve_end; - repeat (destruct i as [|i]; simpl; [|apply lt_S_n in Hi]; try solve_end); clear Hi; - repeat (destruct j as [|j]; simpl; [|apply lt_S_n in Hj]; try solve_end); clear Hj. + repeat (destruct i as [|i]; simpl; [|apply <- Nat.succ_lt_mono in Hi]; try solve_end); clear Hi; + repeat (destruct j as [|j]; simpl; [|apply <- Nat.succ_lt_mono in Hj]; try solve_end); clear Hj. Ltac lma' := apply mat_equiv_eq; @@ -543,10 +543,12 @@ Ltac show_wf := let y := fresh "y" in let H := fresh "H" in intros x y [H | H]; - apply le_plus_minus in H; rewrite H; + apply Nat.sub_add in H; + rewrite Nat.add_comm in H; + rewrite <- H; cbv; destruct_m_eq; - try lca. + try lca. (* Create HintDb wf_db. *) #[export] Hint Resolve WF_Zero WF_I WF_I1 WF_mult WF_plus WF_scale WF_transpose @@ -1191,7 +1193,7 @@ Proof. + repeat rewrite <- beq_nat_refl; simpl. destruct n. - simpl. - rewrite mult_0_r. + rewrite Nat.mul_0_r. bdestruct (y 0) by lia. assert (n * q * s <> 0) by lia. apply Nat.neq_mul_0 in H as [Hmp Hr]. @@ -1312,7 +1314,7 @@ Proof. prep_matrix_equality. destruct q. + simpl. - rewrite mult_0_r. + rewrite Nat.mul_0_r. simpl. rewrite Cmult_0_r. reflexivity. @@ -2622,7 +2624,7 @@ Proof. intros. apply big_sum_eq_bounded; intros. bdestruct (e + S x0 =? e); try lia; easy. unfold get_vec. simpl. - rewrite plus_0_r; easy. + rewrite Nat.add_0_r; easy. Qed. (* shows that we can eliminate a column in a matrix using col_add_many *) @@ -2637,13 +2639,16 @@ Proof. intros. assert (H' : (big_sum (fun x : nat => (as' x 0 .* get_vec x T) i 0) (S m) = (@Mmult n m 1 (reduce_col T col) (reduce_row as' col)) i 0)%C). { unfold Mmult. - replace (S m) with (col + (S (m - col))) by lia; rewrite big_sum_sum. - rewrite (le_plus_minus col m); try lia; rewrite big_sum_sum. + replace (S m) with (col + (S (m - col))) by lia; rewrite big_sum_sum. + rewrite <- (Nat.sub_add col m); + try rewrite (Nat.add_comm (m - col) col); + try lia; rewrite big_sum_sum. apply Cplus_simplify. apply big_sum_eq_bounded; intros. unfold get_vec, scale, reduce_col, reduce_row. bdestruct (x let d := fresh "d" in destruct (lt_ex_diff_r _ _ R); clear R; subst - | H : _ = _ |- _ => rewrite <- plus_assoc in H + | H : _ = _ |- _ => rewrite <- Nat.add_assoc in H | H : ?a + _ = ?a + _ |- _ => apply Nat.add_cancel_l in H; subst | H : ?a + _ = ?b + _ |- _ => destruct (lt_eq_lt_dec a b) as [[?|?]|?]; subst end; try lia. @@ -4375,7 +4387,7 @@ Ltac gridify := restore_dims; distribute_plus; repeat rewrite Nat.pow_add_r; repeat rewrite <- id_kron; simpl; - repeat rewrite mult_assoc; + repeat rewrite Nat.mul_assoc; restore_dims; repeat rewrite <- kron_assoc by auto_wf; restore_dims; repeat rewrite kron_mixed_product; (* simplify *) diff --git a/Measurement.v b/Measurement.v index 7b5b1ec..b92d81a 100644 --- a/Measurement.v +++ b/Measurement.v @@ -59,7 +59,7 @@ Proof. unfold prob_partial_meas. rewrite norm_squared. unfold inner_product, Mmult, adjoint. - rewrite (@big_sum_func_distr C R _ C_is_group _ R_is_group), mult_1_l. + rewrite (@big_sum_func_distr C R _ C_is_group _ R_is_group), Nat.mul_1_l. apply big_sum_eq_bounded; intros. unfold probability_of_outcome. assert (H' : forall c, ((Cmod c)^2)%R = fst (c^* * c)). diff --git a/Pad.v b/Pad.v index 442a79e..c7b8e43 100644 --- a/Pad.v +++ b/Pad.v @@ -219,29 +219,29 @@ Ltac rem_max_min := unfold gt, ge, abs_diff in *; repeat match goal with | [ H: (?a < ?b)%nat |- context[Nat.max (?a - ?b) (?b - ?a)] ] => - rewrite (Max.max_r (a - b) (b - a)) by lia + rewrite (Nat.max_r (a - b) (b - a)) by lia | [ H: (?a < ?b)%nat |- context[Nat.max (?b - ?a) (?a - ?b)] ] => - rewrite (Max.max_l (b - a) (a - b)) by lia + rewrite (Nat.max_l (b - a) (a - b)) by lia | [ H: (?a <= ?b)%nat |- context[Nat.max (?a - ?b) (?b - ?a)] ] => - rewrite (Max.max_r (a - b) (b - a)) by lia + rewrite (Nat.max_r (a - b) (b - a)) by lia | [ H: (?a <= ?b)%nat |- context[Nat.max (?b - ?a) (?a - ?b)] ] => - rewrite (Max.max_l (b - a) (a - b)) by lia + rewrite (Nat.max_l (b - a) (a - b)) by lia | [ H: (?a < ?b)%nat |- context[Nat.min ?a ?b] ] => - rewrite Min.min_l by lia + rewrite Nat.min_l by lia | [ H: (?a < ?b)%nat |- context[Nat.max ?a ?b] ] => - rewrite Max.max_r by lia + rewrite Nat.max_r by lia | [ H: (?a < ?b)%nat |- context[Nat.min ?b ?a] ] => - rewrite Min.min_r by lia + rewrite Nat.min_r by lia | [ H: (?a < ?b)%nat |- context[Nat.max ?b ?a] ] => - rewrite Max.max_l by lia + rewrite Nat.max_l by lia | [ H: (?a <= ?b)%nat |- context[Nat.min ?a ?b] ] => - rewrite Min.min_l by lia + rewrite Nat.min_l by lia | [ H: (?a <= ?b)%nat |- context[Nat.max ?a ?b] ] => - rewrite Max.max_r by lia + rewrite Nat.max_r by lia | [ H: (?a <= ?b)%nat |- context[Nat.min ?b ?a] ] => - rewrite Min.min_r by lia + rewrite Nat.min_r by lia | [ H: (?a <= ?b)%nat |- context[Nat.max ?b ?a] ] => - rewrite Max.max_l by lia + rewrite Nat.max_l by lia end. Definition pad2x2_u (dim n : nat) (u : Square 2) : Square (2^dim) := @embed dim [(u,n)]. @@ -336,7 +336,8 @@ Proof. intros. unfold pad2x2_u, pad2x2_ctrl, abs_diff. bdestruct_all; simpl; rem_max_min; Msimpl; trivial. + repeat rewrite Mmult_plus_distr_l; repeat rewrite Mmult_plus_distr_r. - rewrite le_plus_minus_r by (apply Nat.lt_le_incl; trivial). + rewrite (Nat.add_comm n (o - n)). + rewrite Nat.sub_add by (apply Nat.lt_le_incl; trivial). repeat rewrite embed_mult; simpl. rewrite (embed_commutes dim [(A, m); (∣1⟩⟨1∣, n); (B, o)] [(∣1⟩⟨1∣, n); (B, o); (A, m)]). rewrite (embed_commutes dim [(A, m); (∣0⟩⟨0∣, n); (I 2, o)] [(∣0⟩⟨0∣, n); (I 2, o); (A, m)]). @@ -345,7 +346,8 @@ Proof. all : rewrite perm_swap; apply perm_skip; apply perm_swap. + repeat rewrite Mmult_plus_distr_l; repeat rewrite Mmult_plus_distr_r. - rewrite le_plus_minus_r by trivial; repeat rewrite embed_mult; simpl. + rewrite (Nat.add_comm o (n - o)). + rewrite Nat.sub_add by trivial; repeat rewrite embed_mult; simpl. rewrite (embed_commutes dim [(A, m); (B, o); (∣1⟩⟨1∣, n)] [(B, o); (∣1⟩⟨1∣, n); (A, m)]). rewrite (embed_commutes dim [(A, m); (I 2, o); (∣0⟩⟨0∣, n)] [(I 2, o); (∣0⟩⟨0∣, n); (A, m)]). trivial. diff --git a/Quantum.v b/Quantum.v index ee8fcfd..a2cdad2 100644 --- a/Quantum.v +++ b/Quantum.v @@ -898,7 +898,8 @@ Proof. ++ apply functional_extensionality. intros x. bdestructΩ (n + x (y < 2)%nat -> cnot × ∣ x,y ⟩ = ∣ x, (x + y) mod 2 ⟩. Proof. - intros; destruct x as [| [|x]], y as [| [|y]]; try lia; lma'. -Qed. + intros; destruct x as [| [|x]], y as [| [|y]]; try lia; lma'. Qed. Lemma CNOT00_spec : cnot × ∣ 0,0 ⟩ = ∣ 0,0 ⟩. Proof. lma'. Qed. @@ -646,7 +645,8 @@ Proof. repeat rewrite Nat.add_0_r. rewrite <- Nat.add_succ_l. replace (S (2 ^ n - 1)) with (1 + (2 ^ n - 1)) by easy. - rewrite <- le_plus_minus. + rewrite (Nat.add_comm 1 (2 ^ n - 1)). + rewrite Nat.sub_add. rewrite <- Nat.add_sub_assoc. reflexivity. all: induction n; simpl; try lia. @@ -1605,8 +1605,8 @@ Qed. Lemma product_0 : forall f n, product (fun _ : nat => false) f n = false. Proof. intros f n. - induction n; simpl; auto. - rewrite IHn; reflexivity. + induction n; simpl; auto; + rewrite IHn; reflexivity. Qed. Lemma nat_to_funbool_0 : forall n, nat_to_funbool n 0 = (fun _ => false).