diff --git a/thys/Isabelle_Marries_Dirac/Binary_Nat.thy b/thys/Isabelle_Marries_Dirac/Binary_Nat.thy --- a/thys/Isabelle_Marries_Dirac/Binary_Nat.thy +++ b/thys/Isabelle_Marries_Dirac/Binary_Nat.thy @@ -1,246 +1,246 @@ (* Authors: Anthony Bordg, University of Cambridge, apdb3@cam.ac.uk Hanna Lachnitt, TU Wien, lachnitt@student.tuwien.ac.at *) section \Binary Representation of Natural Numbers\ theory Binary_Nat imports HOL.Nat HOL.List Basics begin primrec bin_rep_aux:: "nat \ nat \ nat list" where "bin_rep_aux 0 m = [m]" | "bin_rep_aux (Suc n) m = m div 2^n # bin_rep_aux n (m mod 2^n)" lemma length_of_bin_rep_aux: fixes n m:: nat assumes "m < 2^n" shows "length (bin_rep_aux n m) = n+1" using assms proof(induction n arbitrary: m) case 0 then show "length (bin_rep_aux 0 m) = 0 + 1" by simp next case (Suc n) assume a0:"\m. m < 2^n \ length (bin_rep_aux n m) = n + 1" and "m < 2^(Suc n)" then show "length (bin_rep_aux (Suc n) m) = Suc n + 1" using a0 by simp qed lemma bin_rep_aux_neq_nil: fixes n m:: nat shows "bin_rep_aux n m \ []" using bin_rep_aux.simps by (metis list.distinct(1) old.nat.exhaust) lemma last_of_bin_rep_aux: fixes n m:: nat assumes "m < 2^n" and "m \ 0" shows "last (bin_rep_aux n m) = 0" using assms proof(induction n arbitrary: m) case 0 assume "m < 2^0" and "m \ 0" then show "last (bin_rep_aux 0 m) = 0" by simp next case (Suc n) assume a0:"\m. m < 2^n \ m \ 0 \ last (bin_rep_aux n m) = 0" and "m < 2^(Suc n)" and "m \ 0" then show "last (bin_rep_aux (Suc n) m) = 0" using bin_rep_aux_neq_nil by simp qed lemma mod_mod_power_cancel: fixes m n p:: nat assumes "m \ n" shows "p mod 2^n mod 2^m = p mod 2^m" using assms by (simp add: dvd_power_le mod_mod_cancel) lemma bin_rep_aux_index: fixes n m i:: nat assumes "n \ 1" and "m < 2^n" and "m \ 0" and "i \ n" shows "bin_rep_aux n m ! i = (m mod 2^(n-i)) div 2^(n-1-i)" using assms proof(induction n arbitrary: m i rule: nat_induct_at_least) case base assume "m < 2^1" and "i \ 1" then show "bin_rep_aux 1 m ! i = m mod 2^(1-i) div 2^(1-1-i)" using bin_rep_aux.simps by (metis One_nat_def base.prems(2) diff_is_0_eq' diff_zero div_by_1 le_Suc_eq le_numeral_extra(3) nth_Cons' power_0 unique_euclidean_semiring_numeral_class.mod_less) next case (Suc n) assume a0:"\m i. m < 2^n \ m \ 0 \ i \ n \ bin_rep_aux n m ! i = m mod 2 ^ (n-i) div 2^(n-1-i)" and a1:"m < 2^(Suc n)" and a2:"i \ Suc n" and a3:"m \ 0" then show "bin_rep_aux (Suc n) m ! i = m mod 2^(Suc n - i) div 2^(Suc n - 1 - i)" proof- have "bin_rep_aux (Suc n) m = m div 2^n # bin_rep_aux n (m mod 2^n)" by simp then have f0:"bin_rep_aux (Suc n) m ! i = (m div 2^n # bin_rep_aux n (m mod 2^n)) ! i" by simp then have "bin_rep_aux (Suc n) m ! i = m div 2^n" if "i = 0" using that by simp then have f1:"bin_rep_aux (Suc n) m ! i = m mod 2^(Suc n - i) div 2^(Suc n - 1 - i)" if "i = 0" proof- have "m mod 2^(Suc n - i) = m" using that a1 by (simp add: Suc.prems(2)) then have "m mod 2^(Suc n - i) div 2^(Suc n - 1 - i) = m div 2^n" using that by simp thus ?thesis by (simp add: that) qed then have "bin_rep_aux (Suc n) m ! i = bin_rep_aux n (m mod 2^n) ! (i-1)" if "i \ 1" using that f0 by simp then have f2:"bin_rep_aux (Suc n) m ! i = ((m mod 2^n) mod 2^(n - (i - 1))) div 2^(n - 1 - (i - 1))" if "i \ 1" using that a0 a1 a2 a3 Suc.prems(2) by simp then have f3:"bin_rep_aux (Suc n) m ! i = ((m mod 2^n) mod 2^(Suc n - i)) div 2^(Suc n - 1 - i)" if "i \ 1" using that by simp then have "bin_rep_aux (Suc n) m ! i = m mod 2^(Suc n - i) div 2^(Suc n - 1 - i)" if "i \ 1" proof- have "Suc n - i \ n" using that by simp then have "m mod 2^n mod 2^(Suc n - i) = m mod 2^(Suc n - i)" using mod_mod_power_cancel[of "Suc n - i" "n" "m"] by simp thus ?thesis using that f3 by simp qed thus ?thesis using f1 f2 using linorder_not_less by blast qed qed lemma bin_rep_aux_coeff: fixes n m i:: nat assumes "m < 2^n" and "i \ n" and "m \ 0" shows "bin_rep_aux n m ! i = 0 \ bin_rep_aux n m ! i = 1" using assms proof(induction n arbitrary: m i) case 0 assume "m < 2^0" and "i \ 0" and "m \ 0" then show "bin_rep_aux 0 m ! i = 0 \ bin_rep_aux 0 m ! i = 1" by simp next case (Suc n) assume a0:"\m i. m < 2 ^ n \ i \ n \ m \ 0 \ bin_rep_aux n m ! i = 0 \ bin_rep_aux n m ! i = 1" and a1:"m < 2^Suc n" and a2:"i \ Suc n" and a3:"m \ 0" then show "bin_rep_aux (Suc n) m ! i = 0 \ bin_rep_aux (Suc n) m ! i = 1" proof- have "bin_rep_aux (Suc n) m ! i = (m div 2^n # bin_rep_aux n (m mod 2^n)) ! i" by simp moreover have "\ = bin_rep_aux n (m mod 2^n) ! (i - 1)" if "i \ 1" using that by simp moreover have "m mod 2^n < 2^n" by simp ultimately have "bin_rep_aux (Suc n) m ! i = 0 \ bin_rep_aux (Suc n) m ! i = 1" if "i\1" using that a0[of "m mod 2^n" "i-1"] a2 by simp moreover have "m div 2^n = 0 \ m div 2^n = 1" using a1 a3 less_mult_imp_div_less by(simp add: less_2_cases) ultimately show ?thesis by (simp add: nth_Cons') qed qed definition bin_rep:: "nat \ nat \ nat list" where "bin_rep n m = butlast (bin_rep_aux n m)" lemma length_of_bin_rep: fixes n m:: nat assumes "m < 2^n" shows "length (bin_rep n m) = n" using assms length_of_bin_rep_aux bin_rep_def by simp lemma bin_rep_coeff: fixes n m i:: nat assumes "m < 2^n" and "i < n" and "m \ 0" shows "bin_rep n m ! i = 0 \ bin_rep n m ! i = 1" using assms bin_rep_def bin_rep_aux_coeff length_of_bin_rep by(simp add: nth_butlast) lemma bin_rep_index: fixes n m i:: nat assumes "n \ 1" and "m < 2^n" and "i < n" and "m \ 0" shows "bin_rep n m ! i = (m mod 2^(n-i)) div 2^(n-1-i)" proof- have "bin_rep n m ! i = bin_rep_aux n m ! i" using bin_rep_def length_of_bin_rep nth_butlast assms(3) by (simp add: nth_butlast assms(2)) thus ?thesis using assms bin_rep_aux_index by simp qed lemma bin_rep_eq: fixes n m:: nat assumes "n \ 1" and "m \ 0" and "m < 2^n" and "m \ 0" shows "m = (\i = m mod 2^(n-i) - m mod 2^(n-i) mod 2^(n-1-i)" by (simp add: minus_mod_eq_div_mult) moreover have "\ = int(m mod 2^(n-i)) - m mod 2^(n-i) mod 2^(n-1-i)" using mod_less_eq_dividend of_nat_diff by blast moreover have "\ = int(m mod 2^(n-i)) - m mod 2^(n-1-i)" using mod_mod_power_cancel[of "n-1-i" "n-i"] by (simp add: dvd_power_le mod_mod_cancel) ultimately have "bin_rep n m ! i * 2^(n-1-i) = int (m mod 2^(n-i)) - m mod 2^(n-1-i)" by presburger } then have f0:"(\iiiiii. (m mod 2^(n-i)))::nat\nat" "(\i. (m mod 2^(n-1-i)))::nat\nat" "{..<(n::nat)}"] by auto moreover have "\ = m mod 2^n + (\i\{1..i = m mod 2^n + (\iii. m mod 2 ^ (n - 1 - i)" "1" "n-1"]) - by (smt One_nat_def assms(1) le_add_diff_inverse lessThan_atLeast0 plus_1_eq_Suc sum.cong sum.shift_bounds_Suc_ivl) + by (smt (verit) One_nat_def assms(1) le_add_diff_inverse lessThan_atLeast0 plus_1_eq_Suc sum.cong sum.shift_bounds_Suc_ivl) moreover have "\ = m mod 2^n - m mod 2^0" by simp moreover have "\ = m" using assms by auto ultimately show "m = (\i n" shows "(bin_rep k m) ! 0 = 0" proof- have "m < 2^(k-1)" - using assms by(smt Suc_diff_1 Suc_leI gr0I le_trans less_or_eq_imp_le linorder_neqE_nat not_less + using assms by (smt (verit) Suc_diff_1 Suc_leI gr0I le_trans less_or_eq_imp_le linorder_neqE_nat not_less one_less_numeral_iff power_strict_increasing semiring_norm(76)) then have f:"m div 2^(k-1) = 0" by auto have "k \ 1" using assms(2) by simp moreover have "bin_rep_aux k m = (m div 2^(k-1)) # (bin_rep_aux (k-1) (m mod 2^(k-1)))" using bin_rep_aux.simps(2) by(metis Suc_diff_1 assms(2) diff_0_eq_0 neq0_conv zero_less_diff) moreover have "bin_rep k m = butlast ((m div 2^(k-1)) # (bin_rep_aux (k-1) (m mod 2^(k-1))))" using bin_rep_def by (simp add: calculation(2)) moreover have "\ = butlast (0 # (bin_rep_aux (k-1) (m mod 2^(k-1))))" using f by simp moreover have "\ = 0 # butlast (bin_rep_aux (k-1) (m mod 2^(k-1)))" by(simp add: bin_rep_aux_neq_nil) ultimately show ?thesis by simp qed lemma bin_rep_index_0_geq: fixes n m:: nat assumes "m \ 2^n" and "m < 2^(n+1)" shows "bin_rep (n+1) m ! 0 = 1" proof- have "bin_rep (Suc n) m = butlast (bin_rep_aux (Suc n) m)" using bin_rep_def by simp moreover have "\ = butlast (1 # (bin_rep_aux n (m mod 2^n)))" using assms bin_rep_aux_def by simp moreover have "\ = 1 # butlast (bin_rep_aux n (m mod 2^n))" by (simp add: bin_rep_aux_neq_nil) ultimately show ?thesis by (simp add: bin_rep_aux_neq_nil) qed end \ No newline at end of file diff --git a/thys/Isabelle_Marries_Dirac/Complex_Vectors.thy b/thys/Isabelle_Marries_Dirac/Complex_Vectors.thy --- a/thys/Isabelle_Marries_Dirac/Complex_Vectors.thy +++ b/thys/Isabelle_Marries_Dirac/Complex_Vectors.thy @@ -1,236 +1,236 @@ (* Author: Anthony Bordg, University of Cambridge, apdb3@cam.ac.uk *) section \Complex Vectors\ theory Complex_Vectors imports Quantum VectorSpace.VectorSpace begin subsection \The Vector Space of Complex Vectors of Dimension n\ definition module_cpx_vec:: "nat \ (complex, complex vec) module" where "module_cpx_vec n \ module_vec TYPE(complex) n" definition cpx_rng:: "complex ring" where "cpx_rng \ \carrier = UNIV, mult = (*), one = 1, zero = 0, add = (+)\" lemma cpx_cring_is_field [simp]: "field cpx_rng" apply unfold_locales apply (auto intro: right_inverse simp: cpx_rng_def Units_def field_simps) by (metis add.right_neutral add_diff_cancel_left' add_uminus_conv_diff) lemma cpx_abelian_monoid [simp]: "abelian_monoid cpx_rng" using cpx_cring_is_field by (simp add: field_def abelian_group_def cring_def domain_def ring_def) lemma vecspace_cpx_vec [simp]: "vectorspace cpx_rng (module_cpx_vec n)" apply unfold_locales apply (auto simp: cpx_rng_def module_cpx_vec_def module_vec_def Units_def field_simps) apply (auto intro: right_inverse add_inv_exists_vec) by (metis add.right_neutral add_diff_cancel_left' add_uminus_conv_diff) lemma module_cpx_vec [simp]: "Module.module cpx_rng (module_cpx_vec n)" using vecspace_cpx_vec by (simp add: vectorspace_def) definition state_basis:: "nat \ nat \ complex vec" where "state_basis n i \ unit_vec (2^n) i" definition unit_vectors:: "nat \ (complex vec) set" where "unit_vectors n \ {unit_vec n i | i::nat. 0 \ i \ i < n}" lemma unit_vectors_carrier_vec [simp]: "unit_vectors n \ carrier_vec n" using unit_vectors_def by auto lemma (in Module.module) finsum_over_singleton [simp]: assumes "f x \ carrier M" shows "finsum M f {x} = f x" using assms by simp lemma lincomb_over_singleton [simp]: assumes "x \ carrier_vec n" and "f \ {x} \ UNIV" shows "module.lincomb (module_cpx_vec n) f {x} = f x \\<^sub>v x" using assms module.lincomb_def module_cpx_vec module_cpx_vec_def module.finsum_over_singleton - by (smt module_vec_simps(3) module_vec_simps(4) smult_carrier_vec) + by (smt (verit) module_vec_simps(3) module_vec_simps(4) smult_carrier_vec) lemma dim_vec_lincomb [simp]: assumes "finite F" and "f: F \ UNIV" and "F \ carrier_vec n" shows "dim_vec (module.lincomb (module_cpx_vec n) f F) = n" using assms proof(induct F) case empty show "dim_vec (module.lincomb (module_cpx_vec n) f {}) = n" proof - have "module.lincomb (module_cpx_vec n) f {} = 0\<^sub>v n" using module.lincomb_def abelian_monoid.finsum_empty module_cpx_vec_def vecspace_cpx_vec vectorspace_def - by (smt abelian_group_def Module.module_def module_vec_simps(2)) + by (smt (verit) abelian_group_def Module.module_def module_vec_simps(2)) thus ?thesis by simp qed next case (insert x F) hence "module.lincomb (module_cpx_vec n) f (insert x F) = (f x \\<^sub>v x) \\<^bsub>module_cpx_vec n\<^esub> module.lincomb (module_cpx_vec n) f F" using module_cpx_vec_def module_vec_def module_cpx_vec module.lincomb_insert cpx_rng_def insert_subset - by (smt Pi_I' UNIV_I Un_insert_right module_vec_simps(4) partial_object.select_convs(1) sup_bot.comm_neutral) + by (smt (verit) Pi_I' UNIV_I Un_insert_right module_vec_simps(4) partial_object.select_convs(1) sup_bot.comm_neutral) hence "dim_vec (module.lincomb (module_cpx_vec n) f (insert x F)) = dim_vec (module.lincomb (module_cpx_vec n) f F)" using index_add_vec by (simp add: module_cpx_vec_def module_vec_simps(1)) thus "dim_vec (module.lincomb (module_cpx_vec n) f (insert x F)) = n" using insert.hyps(3) insert.prems(2) by simp qed lemma lincomb_vec_index [simp]: assumes "finite F" and a2:"i < n" and "F \ carrier_vec n" and "f: F \ UNIV" shows "module.lincomb (module_cpx_vec n) f F $ i = (\v\F. f v * (v $ i))" using assms proof(induct F) case empty then show "module.lincomb (module_cpx_vec n) f {} $ i = (\v\{}. f v * v $ i)" apply auto using a2 module.lincomb_def abelian_monoid.finsum_empty module_cpx_vec_def by (metis (mono_tags) abelian_group_def index_zero_vec(1) module_cpx_vec Module.module_def module_vec_simps(2)) next case(insert x F) have "module.lincomb (module_cpx_vec n) f (insert x F) = f x \\<^sub>v x \\<^bsub>module_cpx_vec n\<^esub> module.lincomb (module_cpx_vec n) f F" using module.lincomb_insert module_cpx_vec insert.hyps(1) module_cpx_vec_def module_vec_def insert.prems(2) insert.hyps(2) insert.prems(3) insert_def - by (smt Pi_I' UNIV_I Un_insert_right cpx_rng_def insert_subset module_vec_simps(4) + by (smt (verit) Pi_I' UNIV_I Un_insert_right cpx_rng_def insert_subset module_vec_simps(4) partial_object.select_convs(1) sup_bot.comm_neutral) then have "module.lincomb (module_cpx_vec n) f (insert x F) $ i = (f x \\<^sub>v x) $ i + module.lincomb (module_cpx_vec n) f F $ i" using index_add_vec(1) a2 dim_vec_lincomb by (metis Pi_split_insert_domain insert.hyps(1) insert.prems(2) insert.prems(3) insert_subset module_cpx_vec_def module_vec_simps(1)) hence "module.lincomb (module_cpx_vec n) f (insert x F) $ i = f x * x $ i + (\v\F. f v * v $ i)" using index_smult_vec a2 insert.prems(2) insert_def insert.hyps(3) by auto with insert show "module.lincomb (module_cpx_vec n) f (insert x F) $ i = (\v\insert x F. f v * v $ i)" by auto qed lemma unit_vectors_is_lin_indpt [simp]: "module.lin_indpt cpx_rng (module_cpx_vec n) (unit_vectors n)" proof assume "module.lin_dep cpx_rng (module_cpx_vec n) (unit_vectors n)" hence "\A a v. (finite A \ A \ (unit_vectors n) \ (a \ A \ UNIV) \ (module.lincomb (module_cpx_vec n) a A = \\<^bsub>module_cpx_vec n\<^esub>) \ (v \ A) \ (a v \ \\<^bsub>cpx_rng\<^esub>))" - using module.lin_dep_def cpx_rng_def module_cpx_vec by (smt Pi_UNIV UNIV_I) + using module.lin_dep_def cpx_rng_def module_cpx_vec by (smt (verit) Pi_UNIV UNIV_I) moreover obtain A and a and v where f1:"finite A" and f2:"A \ (unit_vectors n)" and "a \ A \ UNIV" and f4:"module.lincomb (module_cpx_vec n) a A = \\<^bsub>module_cpx_vec n\<^esub>" and f5:"v \ A" and f6:"a v \ \\<^bsub>cpx_rng\<^esub>" using calculation by blast moreover obtain i where f7:"v = unit_vec n i" and f8:"i < n" using unit_vectors_def calculation by auto ultimately have f9:"module.lincomb (module_cpx_vec n) a A $ i = (\u\A. a u * (u $ i))" using lincomb_vec_index - by (smt carrier_dim_vec index_unit_vec(3) mem_Collect_eq subset_iff sum.cong unit_vectors_def) + by (smt (verit) carrier_dim_vec index_unit_vec(3) mem_Collect_eq subset_iff sum.cong unit_vectors_def) moreover have "\u\A.\j j \ i \ a u * (u $ i) = 0" using unit_vectors_def index_unit_vec by (simp add: f8) then have "(\u\A. a u * (u $ i)) = (\u\A. if u=v then a v * v $ i else 0)" - using f2 unit_vectors_def f7 by (smt mem_Collect_eq subsetCE sum.cong) + using f2 unit_vectors_def f7 by (smt (verit) mem_Collect_eq subsetCE sum.cong) also have "\ = a v * (v $ i)" using abelian_monoid.finsum_singleton[of cpx_rng v A "\u\A. a u * (u $ i)"] cpx_abelian_monoid f5 f1 cpx_rng_def by simp also have "\ = a v" using f7 index_unit_vec f8 by simp also have "\ \ 0" using f6 by (simp add: cpx_rng_def) finally show False using f4 module_cpx_vec_def module_vec_def index_zero_vec f8 f9 by (simp add: module_vec_simps(2)) qed lemma unit_vectors_is_genset [simp]: "module.gen_set cpx_rng (module_cpx_vec n) (unit_vectors n)" proof show "module.span cpx_rng (module_cpx_vec n) (unit_vectors n) \ carrier (module_cpx_vec n)" using module.span_def dim_vec_lincomb carrier_vec_def cpx_rng_def - by (smt Collect_mono index_unit_vec(3) module.span_is_subset2 module_cpx_vec module_cpx_vec_def + by (smt (verit) Collect_mono index_unit_vec(3) module.span_is_subset2 module_cpx_vec module_cpx_vec_def module_vec_simps(3) unit_vectors_def) next show "carrier (module_cpx_vec n) \ module.span cpx_rng (module_cpx_vec n) (unit_vectors n)" proof fix v assume a1:"v \ carrier (module_cpx_vec n)" define A a lc where "A = {unit_vec n i ::complex vec| i::nat. i < n \ v $ i \ 0}" and "a = (\u\A. u \ v)" and "lc = module.lincomb (module_cpx_vec n) a A" then have f1:"finite A" by simp have f2:"A \ carrier_vec n" using carrier_vec_def A_def by auto have f3:"a \ A \ UNIV" using a_def by simp then have f4:"dim_vec v = dim_vec lc" using f1 f2 f3 a1 module_cpx_vec_def dim_vec_lincomb lc_def by (simp add: module_vec_simps(3)) then have f5:"i < n \ lc $ i = (\u\A. u \ v * u $ i)" for i using lincomb_vec_index lc_def a_def f1 f2 f3 by simp then have "i < n \ j < n \ j \ i \ unit_vec n j \ v * unit_vec n j $ i = 0" for i j by simp then have "i < n \ lc $ i = (\u\A. if u = unit_vec n i then v $ i else 0)" for i using a1 A_def f5 scalar_prod_left_unit - by (smt f4 carrier_vecI dim_vec_lincomb f1 f2 f3 index_unit_vec(2) lc_def + by (smt (verit) f4 carrier_vecI dim_vec_lincomb f1 f2 f3 index_unit_vec(2) lc_def mem_Collect_eq mult.right_neutral sum.cong) then have "i < n \ lc $ i = v $ i" for i using abelian_monoid.finsum_singleton[of cpx_rng i] A_def cpx_rng_def by simp then have f6:"v = lc" using eq_vecI f4 dim_vec_lincomb f1 f2 lc_def by auto have "A \ unit_vectors n" using A_def unit_vectors_def by auto thus "v \ module.span cpx_rng (module_cpx_vec n) (unit_vectors n)" using f6 module.span_def[of cpx_rng "module_cpx_vec n"] lc_def f1 f2 cpx_rng_def module_cpx_vec - by (smt Pi_I' UNIV_I mem_Collect_eq partial_object.select_convs(1)) + by (smt (verit) Pi_I' UNIV_I mem_Collect_eq partial_object.select_convs(1)) qed qed lemma unit_vectors_is_basis [simp]: "vectorspace.basis cpx_rng (module_cpx_vec n) (unit_vectors n)" proof - fix n have "unit_vectors n \ carrier (module_cpx_vec n)" using unit_vectors_def module_cpx_vec_def module_vec_simps(3) by fastforce then show ?thesis using vectorspace.basis_def unit_vectors_is_lin_indpt unit_vectors_is_genset vecspace_cpx_vec - by(smt carrier_dim_vec index_unit_vec(3) mem_Collect_eq module_cpx_vec_def module_vec_simps(3) + by (smt (verit) carrier_dim_vec index_unit_vec(3) mem_Collect_eq module_cpx_vec_def module_vec_simps(3) subsetI unit_vectors_def) qed lemma state_qbit_is_lincomb [simp]: "state_qbit n = {module.lincomb (module_cpx_vec (2^n)) a A|a A. finite A \ A\(unit_vectors (2^n)) \ a\ A \ UNIV \ \module.lincomb (module_cpx_vec (2^n)) a A\ = 1}" proof show "state_qbit n \ {module.lincomb (module_cpx_vec (2^n)) a A |a A. finite A \ A \ unit_vectors (2^n) \ a \ A \ UNIV \ \module.lincomb (module_cpx_vec (2^n)) a A\ = 1}" proof fix v assume a1:"v \ state_qbit n" then show "v \ {module.lincomb (module_cpx_vec (2^n)) a A |a A. finite A \ A \ unit_vectors (2^n) \ a \ A \ UNIV \ \module.lincomb (module_cpx_vec (2^n)) a A\ = 1}" proof - obtain a and A where "finite A" and "a\ A \ UNIV" and "A \ unit_vectors (2^n)" and "v = module.lincomb (module_cpx_vec (2^n)) a A" using a1 state_qbit_def unit_vectors_is_basis vectorspace.basis_def module.span_def vecspace_cpx_vec module_cpx_vec module_cpx_vec_def module_vec_def carrier_vec_def - by(smt Pi_UNIV UNIV_I mem_Collect_eq module_vec_simps(3)) + by (smt (verit) Pi_UNIV UNIV_I mem_Collect_eq module_vec_simps(3)) thus ?thesis using a1 state_qbit_def by auto qed qed show "{module.lincomb (module_cpx_vec (2 ^ n)) a A |a A. finite A \ A \ unit_vectors (2 ^ n) \ a \ A \ UNIV \ \module.lincomb (module_cpx_vec (2 ^ n)) a A\ = 1} \ state_qbit n" proof fix v assume "v \ {module.lincomb (module_cpx_vec (2 ^ n)) a A |a A. finite A \ A \ unit_vectors (2 ^ n) \ a \ A \ UNIV \ \module.lincomb (module_cpx_vec (2 ^ n)) a A\ = 1}" then show "v \ state_qbit n" - using state_qbit_def dim_vec_lincomb unit_vectors_carrier_vec by(smt mem_Collect_eq order_trans) + using state_qbit_def dim_vec_lincomb unit_vectors_carrier_vec by (smt (verit) mem_Collect_eq order_trans) qed qed end \ No newline at end of file diff --git a/thys/Isabelle_Marries_Dirac/Deutsch.thy b/thys/Isabelle_Marries_Dirac/Deutsch.thy --- a/thys/Isabelle_Marries_Dirac/Deutsch.thy +++ b/thys/Isabelle_Marries_Dirac/Deutsch.thy @@ -1,526 +1,526 @@ (* Authors: Hanna Lachnitt, TU Wien, lachnitt@student.tuwien.ac.at Anthony Bordg, University of Cambridge, apdb3@cam.ac.uk *) section \The Deutsch Algorithm\ theory Deutsch imports More_Tensor Measurement begin text \ Given a function $f:{0,1}\mapsto {0,1}$, Deutsch's algorithm decides if this function is constant or balanced with a single $f(x)$ circuit to evaluate the function for multiple values of $x$ simultaneously. The algorithm makes use of quantum parallelism and quantum interference. \ text \ A constant function with values in {0,1} returns either always 0 or always 1. A balanced function is 0 for half of the inputs and 1 for the other half. \ locale deutsch = fixes f:: "nat \ nat" assumes dom: "f \ ({0,1} \\<^sub>E {0,1})" context deutsch begin definition is_swap:: bool where "is_swap = (\x \ {0,1}. f x = 1 - x)" lemma is_swap_values: assumes "is_swap" shows "f 0 = 1" and "f 1 = 0" using assms is_swap_def by auto lemma is_swap_sum_mod_2: assumes "is_swap" shows "(f 0 + f 1) mod 2 = 1" using assms is_swap_def by simp definition const:: "nat \ bool" where "const n = (\x \ {0,1}.(f x = n))" definition is_const:: "bool" where "is_const \ const 0 \ const 1" definition is_balanced:: "bool" where "is_balanced \ (\x \ {0,1}.(f x = x)) \ is_swap" lemma f_values: "(f 0 = 0 \ f 0 = 1) \ (f 1 = 0 \ f 1 = 1)" using dom by auto lemma f_cases: shows "is_const \ is_balanced" using dom is_balanced_def const_def is_const_def is_swap_def f_values by auto lemma const_0_sum_mod_2: assumes "const 0" shows "(f 0 + f 1) mod 2 = 0" using assms const_def by simp lemma const_1_sum_mod_2: assumes "const 1" shows "(f 0 + f 1) mod 2 = 0" using assms const_def by simp lemma is_const_sum_mod_2: assumes "is_const" shows "(f 0 + f 1) mod 2 = 0" using assms is_const_def const_0_sum_mod_2 const_1_sum_mod_2 by auto lemma id_sum_mod_2: assumes "f = id" shows "(f 0 + f 1) mod 2 = 1" using assms by simp lemma is_balanced_sum_mod_2: assumes "is_balanced" shows "(f 0 + f 1) mod 2 = 1" using assms is_balanced_def id_sum_mod_2 is_swap_sum_mod_2 by auto lemma f_ge_0: "\ x. (f x \ 0)" by simp end (* context deutsch *) text \The Deutsch's Transform @{text U\<^sub>f}.\ definition (in deutsch) deutsch_transform:: "complex Matrix.mat" ("U\<^sub>f") where "U\<^sub>f \ mat_of_cols_list 4 [[1 - f(0), f(0), 0, 0], [f(0), 1 - f(0), 0, 0], [0, 0, 1 - f(1), f(1)], [0, 0, f(1), 1 - f(1)]]" lemma (in deutsch) deutsch_transform_dim [simp]: shows "dim_row U\<^sub>f = 4" and "dim_col U\<^sub>f = 4" by (auto simp add: deutsch_transform_def mat_of_cols_list_def) lemma (in deutsch) deutsch_transform_coeff_is_zero [simp]: shows "U\<^sub>f $$ (0,2) = 0" and "U\<^sub>f $$ (0,3) = 0" and "U\<^sub>f $$ (1,2) = 0" and "U\<^sub>f $$(1,3) = 0" and "U\<^sub>f $$ (2,0) = 0" and "U\<^sub>f $$(2,1) = 0" and "U\<^sub>f $$ (3,0) = 0" and "U\<^sub>f $$ (3,1) = 0" using deutsch_transform_def by auto lemma (in deutsch) deutsch_transform_coeff [simp]: shows "U\<^sub>f $$ (0,1) = f(0)" and "U\<^sub>f $$ (1,0) = f(0)" and "U\<^sub>f $$(2,3) = f(1)" and "U\<^sub>f $$ (3,2) = f(1)" and "U\<^sub>f $$ (0,0) = 1 - f(0)" and "U\<^sub>f $$(1,1) = 1 - f(0)" and "U\<^sub>f $$ (2,2) = 1 - f(1)" and "U\<^sub>f $$ (3,3) = 1 - f(1)" using deutsch_transform_def by auto abbreviation (in deutsch) V\<^sub>f:: "complex Matrix.mat" where "V\<^sub>f \ Matrix.mat 4 4 (\(i,j). if i=0 \ j=0 then 1 - f(0) else (if i=0 \ j=1 then f(0) else (if i=1 \ j=0 then f(0) else (if i=1 \ j=1 then 1 - f(0) else (if i=2 \ j=2 then 1 - f(1) else (if i=2 \ j=3 then f(1) else (if i=3 \ j=2 then f(1) else (if i=3 \ j=3 then 1 - f(1) else 0))))))))" lemma (in deutsch) deutsch_transform_alt_rep_coeff_is_zero [simp]: shows "V\<^sub>f $$ (0,2) = 0" and "V\<^sub>f $$ (0,3) = 0" and "V\<^sub>f $$ (1,2) = 0" and "V\<^sub>f $$(1,3) = 0" and "V\<^sub>f $$ (2,0) = 0" and "V\<^sub>f $$(2,1) = 0" and "V\<^sub>f $$ (3,0) = 0" and "V\<^sub>f $$ (3,1) = 0" by auto lemma (in deutsch) deutsch_transform_alt_rep_coeff [simp]: shows "V\<^sub>f $$ (0,1) = f(0)" and "V\<^sub>f $$ (1,0) = f(0)" and "V\<^sub>f $$(2,3) = f(1)" and "V\<^sub>f $$ (3,2) = f(1)" and "V\<^sub>f $$ (0,0) = 1 - f(0)" and "V\<^sub>f $$(1,1) = 1 - f(0)" and "V\<^sub>f $$ (2,2) = 1 - f(1)" and "V\<^sub>f $$ (3,3) = 1 - f(1)" by auto lemma (in deutsch) deutsch_transform_alt_rep: shows "U\<^sub>f = V\<^sub>f" proof show c0:"dim_row U\<^sub>f = dim_row V\<^sub>f" by simp show c1:"dim_col U\<^sub>f = dim_col V\<^sub>f" by simp fix i j:: nat assume "i < dim_row V\<^sub>f" and "j < dim_col V\<^sub>f" then have "i < 4" and "j < 4" by auto thus "U\<^sub>f $$ (i,j) = V\<^sub>f $$ (i,j)" - by (smt deutsch_transform_alt_rep_coeff deutsch_transform_alt_rep_coeff_is_zero deutsch_transform_coeff + by (smt (verit) deutsch_transform_alt_rep_coeff deutsch_transform_alt_rep_coeff_is_zero deutsch_transform_coeff deutsch_transform_coeff_is_zero set_4_disj) qed text \@{text U\<^sub>f} is a gate.\ lemma (in deutsch) transpose_of_deutsch_transform: shows "(U\<^sub>f)\<^sup>t = U\<^sub>f" proof show "dim_row (U\<^sub>f\<^sup>t) = dim_row U\<^sub>f" by simp show "dim_col (U\<^sub>f\<^sup>t) = dim_col U\<^sub>f" by simp fix i j:: nat assume "i < dim_row U\<^sub>f" and "j < dim_col U\<^sub>f" thus "U\<^sub>f\<^sup>t $$ (i, j) = U\<^sub>f $$ (i, j)" by (auto simp add: transpose_mat_def) (metis deutsch_transform_coeff(1-4) deutsch_transform_coeff_is_zero set_4_disj) qed lemma (in deutsch) adjoint_of_deutsch_transform: shows "(U\<^sub>f)\<^sup>\ = U\<^sub>f" proof show "dim_row (U\<^sub>f\<^sup>\) = dim_row U\<^sub>f" by simp show "dim_col (U\<^sub>f\<^sup>\) = dim_col U\<^sub>f" by simp fix i j:: nat assume "i < dim_row U\<^sub>f" and "j < dim_col U\<^sub>f" thus "U\<^sub>f\<^sup>\ $$ (i, j) = U\<^sub>f $$ (i, j)" by (auto simp add: dagger_def) (metis complex_cnj_of_nat complex_cnj_zero deutsch_transform_coeff deutsch_transform_coeff_is_zero set_4_disj) qed lemma (in deutsch) deutsch_transform_is_gate: shows "gate 2 U\<^sub>f" proof show "dim_row U\<^sub>f = 2\<^sup>2" by simp show "square_mat U\<^sub>f" by simp show "unitary U\<^sub>f" proof- have "U\<^sub>f * U\<^sub>f = 1\<^sub>m (dim_col U\<^sub>f)" proof show "dim_row (U\<^sub>f * U\<^sub>f) = dim_row (1\<^sub>m (dim_col U\<^sub>f))" by simp next show "dim_col (U\<^sub>f * U\<^sub>f) = dim_col (1\<^sub>m (dim_col U\<^sub>f))" by simp next fix i j:: nat assume "i < dim_row (1\<^sub>m (dim_col U\<^sub>f))" and "j < dim_col (1\<^sub>m (dim_col U\<^sub>f))" then show "(U\<^sub>f * U\<^sub>f) $$ (i,j) = 1\<^sub>m (dim_col U\<^sub>f) $$ (i, j)" apply (auto simp add: deutsch_transform_alt_rep one_mat_def times_mat_def) apply (auto simp: scalar_prod_def) using f_values by auto qed thus ?thesis by (simp add: adjoint_of_deutsch_transform unitary_def) qed qed text \ Two qubits are prepared. The first one in the state $|0\rangle$, the second one in the state $|1\rangle$. \ abbreviation zero where "zero \ unit_vec 2 0" abbreviation one where "one \ unit_vec 2 1" lemma ket_zero_is_state: shows "state 1 |zero\" by (simp add: state_def ket_vec_def cpx_vec_length_def numerals(2)) lemma ket_one_is_state: shows "state 1 |one\" by (simp add: state_def ket_vec_def cpx_vec_length_def numerals(2)) lemma ket_zero_to_mat_of_cols_list [simp]: "|zero\ = mat_of_cols_list 2 [[1, 0]]" by (auto simp add: ket_vec_def mat_of_cols_list_def) lemma ket_one_to_mat_of_cols_list [simp]: "|one\ = mat_of_cols_list 2 [[0, 1]]" apply (auto simp add: ket_vec_def unit_vec_def mat_of_cols_list_def) using less_2_cases by fastforce text \ Applying the Hadamard gate to the state $|0\rangle$ results in the new state @{term "\\<^sub>0\<^sub>0"} = $\dfrac {(|0\rangle + |1\rangle)} {\sqrt 2 }$ \ abbreviation \\<^sub>0\<^sub>0:: "complex Matrix.mat" where "\\<^sub>0\<^sub>0 \ mat_of_cols_list 2 [[1/sqrt(2), 1/sqrt(2)]]" lemma H_on_ket_zero: shows "(H * |zero\) = \\<^sub>0\<^sub>0" proof fix i j:: nat assume "i < dim_row \\<^sub>0\<^sub>0" and "j < dim_col \\<^sub>0\<^sub>0" then have "i \ {0,1} \ j = 0" by (simp add: mat_of_cols_list_def less_2_cases) then show "(H * |zero\) $$ (i,j) = \\<^sub>0\<^sub>0 $$ (i,j)" by (auto simp add: mat_of_cols_list_def times_mat_def scalar_prod_def H_def) next show "dim_row (H * |zero\) = dim_row \\<^sub>0\<^sub>0" by (simp add: H_def mat_of_cols_list_def) show "dim_col (H * |zero\) = dim_col \\<^sub>0\<^sub>0" by (simp add: H_def mat_of_cols_list_def) qed lemma H_on_ket_zero_is_state: shows "state 1 (H * |zero\)" proof show "gate 1 H" using H_is_gate by simp show "state 1 |zero\" using ket_zero_is_state by simp qed text \ Applying the Hadamard gate to the state $|0\rangle$ results in the new state @{text \\<^sub>0\<^sub>1} = $\dfrac {(|0\rangle - |1\rangle)} {\sqrt 2}$. \ abbreviation \\<^sub>0\<^sub>1:: "complex Matrix.mat" where "\\<^sub>0\<^sub>1 \ mat_of_cols_list 2 [[1/sqrt(2), -1/sqrt(2)]]" lemma H_on_ket_one: shows "(H * |one\) = \\<^sub>0\<^sub>1" proof fix i j:: nat assume "i < dim_row \\<^sub>0\<^sub>1" and "j < dim_col \\<^sub>0\<^sub>1" then have "i \ {0,1} \ j = 0" by (simp add: mat_of_cols_list_def less_2_cases) then show "(H * |one\) $$ (i,j) = \\<^sub>0\<^sub>1 $$ (i,j)" by (auto simp add: mat_of_cols_list_def times_mat_def scalar_prod_def H_def ket_vec_def) next show "dim_row (H * |one\) = dim_row \\<^sub>0\<^sub>1" by (simp add: H_def mat_of_cols_list_def) show "dim_col (H * |one\) = dim_col \\<^sub>0\<^sub>1" by (simp add: H_def mat_of_cols_list_def ket_vec_def) qed lemma H_on_ket_one_is_state: shows "state 1 (H * |one\)" using H_is_gate ket_one_is_state by simp text\ Then, the state @{text \\<^sub>1} = $\dfrac {(|00\rangle - |01\rangle + |10\rangle - |11\rangle)} {2} $ is obtained by taking the tensor product of the states @{text \\<^sub>0\<^sub>0} = $\dfrac {(|0\rangle + |1\rangle)} {\sqrt 2} $ and @{text \\<^sub>0\<^sub>1} = $\dfrac {(|0\rangle - |1\rangle)} {\sqrt 2} $. \ abbreviation \\<^sub>1:: "complex Matrix.mat" where "\\<^sub>1 \ mat_of_cols_list 4 [[1/2, -1/2, 1/2, -1/2]]" lemma \\<^sub>0_to_\\<^sub>1: shows "(\\<^sub>0\<^sub>0 \ \\<^sub>0\<^sub>1) = \\<^sub>1" proof fix i j:: nat assume "i < dim_row \\<^sub>1" and "j < dim_col \\<^sub>1" then have "i \ {0,1,2,3}" and "j = 0" using mat_of_cols_list_def by auto moreover have "complex_of_real (sqrt 2) * complex_of_real (sqrt 2) = 2" by (metis mult_2_right numeral_Bit0 of_real_mult of_real_numeral real_sqrt_four real_sqrt_mult) ultimately show "(\\<^sub>0\<^sub>0 \ \\<^sub>0\<^sub>1) $$ (i,j) = \\<^sub>1 $$ (i,j)" using mat_of_cols_list_def by auto next show "dim_row (\\<^sub>0\<^sub>0 \ \\<^sub>0\<^sub>1) = dim_row \\<^sub>1" by (simp add: mat_of_cols_list_def) show "dim_col (\\<^sub>0\<^sub>0 \ \\<^sub>0\<^sub>1) = dim_col \\<^sub>1" by (simp add: mat_of_cols_list_def) qed lemma \\<^sub>1_is_state: shows "state 2 \\<^sub>1" proof show "dim_col \\<^sub>1 = 1" by (simp add: Tensor.mat_of_cols_list_def) show "dim_row \\<^sub>1 = 2\<^sup>2" by (simp add: Tensor.mat_of_cols_list_def) show "\Matrix.col \\<^sub>1 0\ = 1" using H_on_ket_one_is_state H_on_ket_zero_is_state state.is_normal tensor_state2 \\<^sub>0_to_\\<^sub>1 H_on_ket_one H_on_ket_zero by force qed text \ Next, the gate @{text U\<^sub>f} is applied to the state @{text \\<^sub>1} = $\dfrac {(|00\rangle - |01\rangle + |10\rangle - |11\rangle)} {2}$ and @{text \\<^sub>2}= $\dfrac {(|0f(0)\oplus 0\rangle - |0 f(0) \oplus 1\rangle + |1 f(1)\oplus 0\rangle - |1f(1)\oplus 1\rangle)} {2}$ is obtained. This simplifies to @{text \\<^sub>2}= $\dfrac {(|0f(0)\rangle - |0 \overline{f(0)} \rangle + |1 f(1)\rangle - |1\overline{f(1)}\rangle)} {2}$ \ abbreviation (in deutsch) \\<^sub>2:: "complex Matrix.mat" where "\\<^sub>2 \ mat_of_cols_list 4 [[(1 - f(0))/2 - f(0)/2, f(0)/2 - (1 - f(0))/2, (1 - f(1))/2 - f(1)/2, f(1)/2 - (1- f(1))/2]]" lemma (in deutsch) \\<^sub>1_to_\\<^sub>2: shows "U\<^sub>f * \\<^sub>1 = \\<^sub>2" proof fix i j:: nat assume "i < dim_row \\<^sub>2 " and "j < dim_col \\<^sub>2" then have asm:"i \ {0,1,2,3} \ j = 0 " by (auto simp add: mat_of_cols_list_def) then have "i < dim_row U\<^sub>f \ j < dim_col \\<^sub>1" using deutsch_transform_def mat_of_cols_list_def by auto then have "(U\<^sub>f * \\<^sub>1) $$ (i, j) = (\ k \ {0 ..< dim_vec \\<^sub>1}. (Matrix.row U\<^sub>f i) $ k * (Matrix.col \\<^sub>1 j) $ k)" apply (auto simp add: times_mat_def scalar_prod_def). thus "(U\<^sub>f * \\<^sub>1) $$ (i, j) = \\<^sub>2 $$ (i, j)" using mat_of_cols_list_def deutsch_transform_def asm by auto next show "dim_row (U\<^sub>f * \\<^sub>1) = dim_row \\<^sub>2" by (simp add: mat_of_cols_list_def) show "dim_col (U\<^sub>f * \\<^sub>1) = dim_col \\<^sub>2" by (simp add: mat_of_cols_list_def) qed lemma (in deutsch) \\<^sub>2_is_state: shows "state 2 \\<^sub>2" proof show "dim_col \\<^sub>2 = 1" by (simp add: Tensor.mat_of_cols_list_def) show "dim_row \\<^sub>2 = 2\<^sup>2" by (simp add: Tensor.mat_of_cols_list_def) show "\Matrix.col \\<^sub>2 0\ = 1" using gate_on_state_is_state \\<^sub>1_is_state deutsch_transform_is_gate \\<^sub>1_to_\\<^sub>2 state_def by (metis (no_types, lifting)) qed lemma H_tensor_Id_1: defines d:"v \ mat_of_cols_list 4 [[1/sqrt(2), 0, 1/sqrt(2), 0], [0, 1/sqrt(2), 0, 1/sqrt(2)], [1/sqrt(2), 0, -1/sqrt(2), 0], [0, 1/sqrt(2), 0, -1/sqrt(2)]]" shows "(H \ Id 1) = v" proof show "dim_col (H \ Id 1) = dim_col v" by (simp add: d H_def Id_def mat_of_cols_list_def) show "dim_row (H \ Id 1) = dim_row v" by (simp add: d H_def Id_def mat_of_cols_list_def) fix i j:: nat assume "i < dim_row v" and "j < dim_col v" then have "i \ {0..<4} \ j \ {0..<4}" by (auto simp add: d mat_of_cols_list_def) thus "(H \ Id 1) $$ (i, j) = v $$ (i, j)" by (auto simp add: d Id_def H_def mat_of_cols_list_def) qed lemma H_tensor_Id_1_is_gate: shows "gate 2 (H \ Id 1)" proof show "dim_row (H \ Quantum.Id 1) = 2\<^sup>2" using H_tensor_Id_1 by (simp add: mat_of_cols_list_def) show "square_mat (H \ Quantum.Id 1)" using H_is_gate id_is_gate tensor_gate_sqr_mat by blast show "unitary (H \ Quantum.Id 1)" using H_is_gate gate_def id_is_gate tensor_gate by blast qed text \ Applying the Hadamard gate to the first qubit of @{text \\<^sub>2} results in @{text \\<^sub>3} = $\pm |f(0)\oplus f(1)\rangle \left[ \dfrac {(|0\rangle - |1\rangle)} {\sqrt 2}\right]$ \ abbreviation (in deutsch) \\<^sub>3:: "complex Matrix.mat" where "\\<^sub>3 \ mat_of_cols_list 4 [[(1-f(0))/(2*sqrt(2)) - f(0)/(2*sqrt(2)) + (1-f(1))/(2*sqrt(2)) - f(1)/(2*sqrt(2)), f(0)/(2*sqrt(2)) - (1-f(0))/(2*sqrt(2)) + (f(1)/(2*sqrt(2)) - (1-f(1))/(2*sqrt(2))), (1-f(0))/(2*sqrt(2)) - f(0)/(2*sqrt(2)) - (1-f(1))/(2*sqrt(2)) + f(1)/(2*sqrt(2)), f(0)/(2*sqrt(2)) - (1-f(0))/(2*sqrt(2)) - f(1)/(2*sqrt(2)) + (1-f(1))/(2*sqrt(2))]]" lemma (in deutsch) \\<^sub>2_to_\\<^sub>3: shows "(H \ Id 1) * \\<^sub>2 = \\<^sub>3" proof fix i j:: nat assume "i < dim_row \\<^sub>3" and "j < dim_col \\<^sub>3" then have a0:"i \ {0,1,2,3} \ j = 0" by (auto simp add: mat_of_cols_list_def) then have "i < dim_row (H \ Id 1) \ j < dim_col \\<^sub>2" using mat_of_cols_list_def H_tensor_Id_1 by auto then have "((H \ Id 1)*\\<^sub>2) $$ (i,j) = (\ k \ {0 ..< dim_vec \\<^sub>2}. (Matrix.row (H \ Id 1) i) $ k * (Matrix.col \\<^sub>2 j) $ k)" by (auto simp: times_mat_def scalar_prod_def) thus "((H \ Id 1) * \\<^sub>2) $$ (i, j) = \\<^sub>3 $$ (i, j)" using mat_of_cols_list_def H_tensor_Id_1 a0 f_ge_0 by (auto simp: diff_divide_distrib) next show "dim_row ((H \ Id 1) * \\<^sub>2) = dim_row \\<^sub>3" using H_tensor_Id_1 mat_of_cols_list_def by simp show "dim_col ((H \ Id 1) * \\<^sub>2) = dim_col \\<^sub>3" using H_tensor_Id_1 mat_of_cols_list_def by simp qed lemma (in deutsch) \\<^sub>3_is_state: shows "state 2 \\<^sub>3" proof- have "gate 2 (H \ Id 1)" using H_tensor_Id_1_is_gate by simp thus "state 2 \\<^sub>3" using \\<^sub>2_is_state \\<^sub>2_to_\\<^sub>3 by (metis gate_on_state_is_state) qed text \ Finally, all steps are put together. The result depends on the function f. If f is constant the first qubit of $\pm |f(0)\oplus f(1)\rangle \left[ \dfrac {(|0\rangle - |1\rangle)} {\sqrt 2}\right]$ is 0, if it is is\_balanced it is 1. The algorithm only uses one evaluation of f(x) and will always succeed. \ definition (in deutsch) deutsch_algo:: "complex Matrix.mat" where "deutsch_algo \ (H \ Id 1) * (U\<^sub>f * ((H * |zero\) \ (H * |one\)))" lemma (in deutsch) deutsch_algo_result [simp]: shows "deutsch_algo = \\<^sub>3" using deutsch_algo_def H_on_ket_zero H_on_ket_one \\<^sub>0_to_\\<^sub>1 \\<^sub>1_to_\\<^sub>2 \\<^sub>2_to_\\<^sub>3 by simp lemma (in deutsch) deutsch_algo_result_is_state: shows "state 2 deutsch_algo" using \\<^sub>3_is_state by simp text \ If the function is constant then the measurement of the first qubit should result in the state $|0\rangle$ with probability 1. \ lemma (in deutsch) prob0_deutsch_algo_const: assumes "is_const" shows "prob0 2 deutsch_algo 0 = 1" proof - have "{k| k::nat. (k < 4) \ \ select_index 2 0 k} = {0,1}" using select_index_def by auto then have "prob0 2 deutsch_algo 0 = (\j\{0,1}. (cmod(deutsch_algo $$ (j,0)))\<^sup>2)" using deutsch_algo_result_is_state prob0_def by simp thus "prob0 2 deutsch_algo 0 = 1" using assms is_const_def const_def by auto qed lemma (in deutsch) prob1_deutsch_algo_const: assumes "is_const" shows "prob1 2 deutsch_algo 0 = 0" using assms prob0_deutsch_algo_const prob_sum_is_one[of "2" "deutsch_algo" "0"] deutsch_algo_result_is_state by simp text \ If the function is balanced the measurement of the first qubit should result in the state $|1\rangle$ with probability 1. \ lemma (in deutsch) prob0_deutsch_algo_balanced: assumes "is_balanced" shows "prob0 2 deutsch_algo 0 = 0" proof- have "{k| k::nat. (k < 4) \ \ select_index 2 0 k} = {0,1}" using select_index_def by auto then have "prob0 2 deutsch_algo 0 = (\j \ {0,1}. (cmod(deutsch_algo $$ (j,0)))\<^sup>2)" using deutsch_algo_result_is_state prob0_def by simp thus "prob0 2 deutsch_algo 0 = 0" using is_swap_values assms is_balanced_def by auto qed lemma (in deutsch) prob1_deutsch_algo_balanced: assumes "is_balanced" shows "prob1 2 deutsch_algo 0 = 1" using assms prob0_deutsch_algo_balanced prob_sum_is_one[of "2" "deutsch_algo" "0"] deutsch_algo_result_is_state by simp text \Eventually, the measurement of the first qubit results in $f(0)\oplus f(1)$. \ definition (in deutsch) deutsch_algo_eval:: "real" where "deutsch_algo_eval \ prob1 2 deutsch_algo 0" lemma (in deutsch) sum_mod_2_cases: shows "(f 0 + f 1) mod 2 = 0 \ is_const" and "(f 0 + f 1) mod 2 = 1 \ is_balanced" using f_cases is_balanced_sum_mod_2 is_const_sum_mod_2 by auto lemma (in deutsch) deutsch_algo_eval_is_sum_mod_2: shows "deutsch_algo_eval = (f 0 + f 1) mod 2" using deutsch_algo_eval_def f_cases is_const_sum_mod_2 is_balanced_sum_mod_2 prob1_deutsch_algo_const prob1_deutsch_algo_balanced by auto text \ If the algorithm returns 0 then one concludes that the input function is constant and if it returns 1 then the function is balanced. \ theorem (in deutsch) deutsch_algo_is_correct: shows "deutsch_algo_eval = 0 \ is_const" and "deutsch_algo_eval = 1 \ is_balanced" using deutsch_algo_eval_is_sum_mod_2 sum_mod_2_cases by auto end \ No newline at end of file diff --git a/thys/Isabelle_Marries_Dirac/Deutsch_Jozsa.thy b/thys/Isabelle_Marries_Dirac/Deutsch_Jozsa.thy --- a/thys/Isabelle_Marries_Dirac/Deutsch_Jozsa.thy +++ b/thys/Isabelle_Marries_Dirac/Deutsch_Jozsa.thy @@ -1,1765 +1,1765 @@ (* Authors: Hanna Lachnitt, TU Wien, lachnitt@student.tuwien.ac.at Anthony Bordg, University of Cambridge, apdb3@cam.ac.uk *) section \The Deutsch-Jozsa Algorithm\ theory Deutsch_Jozsa imports Deutsch More_Tensor Binary_Nat begin text \ Given a function $f:{0,1}^n \mapsto {0,1}$, the Deutsch-Jozsa algorithm decides if this function is constant or balanced with a single $f(x)$ circuit to evaluate the function for multiple values of $x$ simultaneously. The algorithm makes use of quantum parallelism and quantum interference. A constant function with values in {0,1} returns either always 0 or always 1. A balanced function is 0 for half of the inputs and 1 for the other half. \ locale bob_fun = fixes f:: "nat \ nat" and n:: "nat" assumes dom: "f \ ({(i::nat). i < 2^n} \\<^sub>E {0,1})" assumes dim: "n \ 1" context bob_fun begin definition const:: "nat \ bool" where "const c = (\x\{i::nat. i<2^n}. f x = c)" definition is_const:: bool where "is_const \ const 0 \ const 1" definition is_balanced:: bool where "is_balanced \ \A B ::nat set. A \ {i::nat. i < 2^n} \ B \ {i::nat. i < 2^n} \ card A = 2^(n-1) \ card B = 2^(n-1) \ (\x\A. f x = 0) \ (\x\B. f x = 1)" lemma is_balanced_inter: fixes A B:: "nat set" assumes "\x \ A. f x = 0" and "\x \ B. f x = 1" shows "A \ B = {}" using assms by auto lemma is_balanced_union: fixes A B:: "nat set" assumes "A \ {i::nat. i < 2^n}" and "B \ {i::nat. i < 2^n}" and "card A = 2^(n-1)" and "card B = 2^(n-1)" and "A \ B = {}" shows "A \ B = {i::nat. i < 2^n}" proof- have "finite A" and "finite B" by (simp add: assms(3) card_ge_0_finite) (simp add: assms(4) card_ge_0_finite) then have "card(A \ B) = 2 * 2^(n-1)" using assms(3-5) by (simp add: card_Un_disjoint) then have "card(A \ B) = 2^n" by (metis Nat.nat.simps(3) One_nat_def dim le_0_eq power_eq_if) moreover have "\ = card({i::nat. i < 2^n})" by simp moreover have "A \ B \ {i::nat. i < 2^n}" using assms(1,2) by simp moreover have "finite ({i::nat. i < 2^n})" by simp ultimately show ?thesis using card_subset_eq[of "{i::nat. i < 2^n}" "A \ B"] by simp qed lemma f_ge_0: "\x. f x \ 0" by simp lemma f_dom_not_zero: shows "f \ ({i::nat. n \ 1 \ i < 2^n} \\<^sub>E {0,1})" using dim dom by simp lemma f_values: "\x \ {(i::nat). i < 2^n} . f x = 0 \ f x = 1" using dom by auto end (* bob_fun *) text \The input function has to be constant or balanced.\ locale jozsa = bob_fun + assumes const_or_balanced: "is_const \ is_balanced " text \ Introduce two customised rules: disjunctions with four disjuncts and induction starting from one instead of zero. \ (* To deal with Uf it is often necessary to do a case distinction with four different cases.*) lemma disj_four_cases: assumes "A \ B \ C \ D" and "A \ P" and "B \ P" and "C \ P" and "D \ P" shows "P" using assms by auto text \The unitary transform @{term U\<^sub>f}.\ definition (in jozsa) jozsa_transform:: "complex Matrix.mat" ("U\<^sub>f") where "U\<^sub>f \ Matrix.mat (2^(n+1)) (2^(n+1)) (\(i,j). if i = j then (1-f(i div 2)) else if i = j + 1 \ odd i then f(i div 2) else if i = j - 1 \ even i \ j\1 then f(i div 2) else 0)" lemma (in jozsa) jozsa_transform_dim [simp]: shows "dim_row U\<^sub>f = 2^(n+1)" and "dim_col U\<^sub>f = 2^(n+1)" by (auto simp add: jozsa_transform_def) lemma (in jozsa) jozsa_transform_coeff_is_zero [simp]: assumes "i < dim_row U\<^sub>f \ j < dim_col U\<^sub>f" shows "(i\j \ \(i=j+1 \ odd i) \ \ (i=j-1 \ even i \ j\1)) \ U\<^sub>f $$ (i,j) = 0" using jozsa_transform_def assms by auto lemma (in jozsa) jozsa_transform_coeff [simp]: assumes "i < dim_row U\<^sub>f \ j < dim_col U\<^sub>f" shows "i = j \ U\<^sub>f $$ (i,j) = 1 - f (i div 2)" and "i = j + 1 \ odd i \ U\<^sub>f $$ (i,j) = f (i div 2)" and "j \ 1 \ i = j - 1 \ even i \ U\<^sub>f $$ (i,j) = f (i div 2)" using jozsa_transform_def assms by auto lemma (in jozsa) U\<^sub>f_mult_without_empty_summands_sum_even: fixes i j A assumes "i < dim_row U\<^sub>f" and "j < dim_col A" and "even i" and "dim_col U\<^sub>f = dim_row A" shows "(\k\{0..< dim_row A}. U\<^sub>f $$ (i,k) * A $$ (k,j)) =(\k\{i,i+1}. U\<^sub>f $$ (i,k) * A $$ (k,j))" proof- have "(\k \ {0..< 2^(n+1)}. U\<^sub>f $$ (i,k) * A $$ (k,j)) = (\k \ {0..f $$ (i,k) * A $$ (k,j)) + (\k \ {i,i+1}. U\<^sub>f $$ (i,k) * A $$ (k,j)) + (\k \ {(i+2)..< 2^(n+1)}. U\<^sub>f $$ (i,k) * A $$ (k,j))" proof- have "{0..< 2^(n+1)} = {0.. {i..< 2^(n+1)} \ {i..< 2^(n+1)} = {i,i+1} \ {(i+2)..<2^(n+1)}" using assms(1-3) by auto moreover have "{0.. {i,i+1} = {} \ {i,i+1} \ {(i+2)..< 2^(n+1)} = {} \ {0.. {(i+2)..< 2^(n+1)} = {}" using assms by simp ultimately show ?thesis using sum.union_disjoint by (metis (no_types, lifting) finite_Un finite_atLeastLessThan is_num_normalize(1) ivl_disj_int_two(3)) qed moreover have "(\k \ {0..f $$ (i,k) * A $$ (k,j)) = 0" proof- have "k \ {0.. (i\k \ \(i=k+1 \ odd i) \ \ (i=k-1 \ even i \ k\1))" for k using assms by auto then have "k \ {0.. U\<^sub>f $$ (i,k) = 0" for k using assms(1) by auto then show ?thesis by simp qed moreover have "(\k \ {(i+2)..< 2^(n+1)}. U\<^sub>f $$ (i,k) * A $$ (k,j)) = 0" proof- have "k\{(i+2)..< 2^(n+1)} \ (i\k \ \(i=k+1 \ odd i) \ \ (i=k-1 \ even i \ k\1))" for k by auto then have "k \ {(i+2)..< 2^(n+1)}\ U\<^sub>f $$ (i,k) = 0" for k using assms(1) by auto then show ?thesis by simp qed moreover have "dim_row A = 2^(n+1)" using assms(4) by simp ultimately show "?thesis" by(metis (no_types, lifting) add.left_neutral add.right_neutral) qed lemma (in jozsa) U\<^sub>f_mult_without_empty_summands_even: fixes i j A assumes "i < dim_row U\<^sub>f" and "j < dim_col A" and "even i" and "dim_col U\<^sub>f = dim_row A" shows "(U\<^sub>f * A) $$ (i,j) = (\k \ {i,i+1}. U\<^sub>f $$ (i,k) * A $$ (k,j))" proof- have "(U\<^sub>f * A) $$ (i,j) = (\ k\{0..< dim_row A}. (U\<^sub>f $$ (i,k)) * (A $$ (k,j)))" using assms(1,2,4) index_matrix_prod by (simp add: atLeast0LessThan) then show ?thesis using assms U\<^sub>f_mult_without_empty_summands_sum_even by simp qed lemma (in jozsa) U\<^sub>f_mult_without_empty_summands_sum_odd: fixes i j A assumes "i < dim_row U\<^sub>f" and "j < dim_col A" and "odd i" and "dim_col U\<^sub>f = dim_row A" shows "(\k\{0..< dim_row A}. U\<^sub>f $$ (i,k) * A $$ (k,j)) =(\k\{i-1,i}. U\<^sub>f $$ (i,k) * A $$ (k,j))" proof- have "(\k\{0..< 2^(n+1)}. U\<^sub>f $$ (i,k) * A $$ (k,j)) = (\k \ {0..f $$ (i,k) * A $$ (k,j)) + (\k \ {i-1,i}. U\<^sub>f $$ (i,k) * A $$ (k,j)) + (\k \ {i+1..< 2^(n+1)}. U\<^sub>f $$ (i,k) * A $$ (k,j))" proof- have "{0..< 2^(n+1)} = {0.. {i-1..< 2^(n+1)} \ {i-1..< 2^(n+1)} = {i-1,i} \ {i+1..<2^(n+1)}" using assms(1-3) by auto moreover have "{0.. {i-1,i} = {} \ {i-1,i} \ {i+1..< 2^(n+1)} = {} \ {0.. {i+1..< 2^(n+1)} = {}" using assms by simp ultimately show ?thesis using sum.union_disjoint by(metis (no_types, lifting) finite_Un finite_atLeastLessThan is_num_normalize(1) ivl_disj_int_two(3)) qed moreover have "(\k \ {0..f $$ (i,k) * A $$ (k,j)) = 0" proof- have "k \ {0.. (i\k \ \(i=k+1 \ odd i) \ \ (i=k-1 \ even i \ k\1))" for k by auto then have "k \ {0.. U\<^sub>f $$ (i,k) = 0" for k using assms(1) by auto then show ?thesis by simp qed moreover have "(\k \ {i+1..< 2^(n+1)}. U\<^sub>f $$ (i,k) * A $$ (k,j)) = 0" using assms(3) by auto moreover have "dim_row A = 2^(n+1)" using assms(4) by simp ultimately show "?thesis" by(metis (no_types, lifting) add.left_neutral add.right_neutral) qed lemma (in jozsa) U\<^sub>f_mult_without_empty_summands_odd: fixes i j A assumes "i < dim_row U\<^sub>f" and "j < dim_col A" and "odd i" and "dim_col U\<^sub>f = dim_row A" shows "(U\<^sub>f * A) $$ (i,j) = (\k \ {i-1,i}. U\<^sub>f $$ (i,k) * A $$ (k,j)) " proof- have "(U\<^sub>f * A) $$ (i,j) = (\k \ {0 ..< dim_row A}. (U\<^sub>f $$ (i,k)) * (A $$ (k,j)))" using assms(1,2,4) index_matrix_prod by (simp add: atLeast0LessThan) then show "?thesis" using assms U\<^sub>f_mult_without_empty_summands_sum_odd by auto qed text \@{term U\<^sub>f} is a gate.\ lemma (in jozsa) transpose_of_jozsa_transform: shows "(U\<^sub>f)\<^sup>t = U\<^sub>f" proof show "dim_row (U\<^sub>f\<^sup>t) = dim_row U\<^sub>f" by simp next show "dim_col (U\<^sub>f\<^sup>t) = dim_col U\<^sub>f" by simp next fix i j:: nat assume a0: "i < dim_row U\<^sub>f" and a1: "j < dim_col U\<^sub>f" then show "U\<^sub>f\<^sup>t $$ (i, j) = U\<^sub>f $$ (i, j)" proof (induct rule: disj_four_cases) show "i=j \ (i=j+1 \ odd i) \ (i=j-1 \ even i \ j\1) \ (i\j \ \(i=j+1 \ odd i) \ \ (i=j-1 \ even i \ j\1))" by linarith next assume "i = j" then show "U\<^sub>f\<^sup>t $$ (i,j) = U\<^sub>f $$ (i,j)" using a0 by simp next assume "(i=j+1 \ odd i)" then show "U\<^sub>f\<^sup>t $$ (i,j) = U\<^sub>f $$ (i,j)" using transpose_mat_def a0 a1 by auto next assume a2:"(i=j-1 \ even i \ j\1)" then have "U\<^sub>f $$ (i,j) = f (i div 2)" using a0 a1 jozsa_transform_coeff by auto moreover have "U\<^sub>f $$ (j,i) = f (i div 2)" using a0 a1 a2 jozsa_transform_coeff by (metis add_diff_assoc2 diff_add_inverse2 even_plus_one_iff even_succ_div_two jozsa_transform_dim) ultimately show "?thesis" using transpose_mat_def a0 a1 by simp next assume a2:"(i\j \ \(i=j+1 \ odd i) \ \ (i=j-1 \ even i \ j\1))" then have "(j\i \ \(j=i+1 \ odd j) \ \ (j=i-1 \ even j \ i\1))" by (metis le_imp_diff_is_add diff_add_inverse even_plus_one_iff le_add1) then have "U\<^sub>f $$ (j,i) = 0" using jozsa_transform_coeff_is_zero a0 a1 by auto moreover have "U\<^sub>f $$ (i,j) = 0" using jozsa_transform_coeff_is_zero a0 a1 a2 by auto ultimately show "U\<^sub>f\<^sup>t $$ (i,j) = U\<^sub>f $$ (i,j)" using transpose_mat_def a0 a1 by simp qed qed lemma (in jozsa) adjoint_of_jozsa_transform: shows "(U\<^sub>f)\<^sup>\ = U\<^sub>f" proof show "dim_row (U\<^sub>f\<^sup>\) = dim_row U\<^sub>f" by simp next show "dim_col (U\<^sub>f\<^sup>\) = dim_col U\<^sub>f" by simp next fix i j:: nat assume a0: "i < dim_row U\<^sub>f" and a1: "j < dim_col U\<^sub>f" then show "U\<^sub>f\<^sup>\ $$ (i,j) = U\<^sub>f $$ (i,j)" proof (induct rule: disj_four_cases) show "i=j \ (i=j+1 \ odd i) \ (i=j-1 \ even i \ j\1) \ (i\j \ \(i=j+1 \ odd i) \ \ (i=j-1 \ even i \ j\1))" by linarith next assume "i=j" then show "U\<^sub>f\<^sup>\ $$ (i,j) = U\<^sub>f $$ (i,j)" using a0 dagger_def by simp next assume "(i=j+1 \ odd i)" then show "U\<^sub>f\<^sup>\ $$ (i,j) = U\<^sub>f $$ (i,j)" using a0 dagger_def by auto next assume a2:"(i=j-1 \ even i \ j\1)" then have "U\<^sub>f $$ (i,j) = f (i div 2)" using a0 a1 jozsa_transform_coeff by auto moreover have "U\<^sub>f\<^sup>\ $$ (j,i) = f (i div 2)" using a1 a2 jozsa_transform_coeff dagger_def by auto ultimately show "U\<^sub>f\<^sup>\ $$ (i,j) = U\<^sub>f $$ (i,j)" by(metis a0 a1 cnj_transpose_is_dagger dim_row_of_dagger index_transpose_mat dagger_of_transpose_is_cnj transpose_of_jozsa_transform) next assume a2: "(i\j \ \(i=j+1 \ odd i) \ \ (i=j-1 \ even i \ j\1))" then have f0:"(i\j \ \(j=i+1 \ odd j) \ \ (j=i-1 \ even j \ i\1))" by (metis le_imp_diff_is_add diff_add_inverse even_plus_one_iff le_add1) then have "U\<^sub>f $$ (j,i) = 0" and "cnj 0 = 0" using jozsa_transform_coeff_is_zero a0 a1 a2 by auto then have "U\<^sub>f\<^sup>\ $$ (i,j) = 0" using a0 a1 dagger_def by simp then show "U\<^sub>f\<^sup>\ $$ (i, j) = U\<^sub>f $$ (i, j)" using a0 a1 a2 jozsa_transform_coeff_is_zero by auto qed qed lemma (in jozsa) jozsa_transform_is_unitary_index_even: fixes i j:: nat assumes "i < dim_row U\<^sub>f" and "j < dim_col U\<^sub>f" and "even i" shows "(U\<^sub>f * U\<^sub>f) $$ (i,j) = 1\<^sub>m (dim_col U\<^sub>f) $$ (i,j)" proof- have "(U\<^sub>f * U\<^sub>f) $$ (i,j) = (\k \ {i,i+1}. U\<^sub>f $$ (i,k) * U\<^sub>f $$ (k,j)) " using U\<^sub>f_mult_without_empty_summands_even[of i j U\<^sub>f ] assms by simp moreover have "U\<^sub>f $$ (i,i) * U\<^sub>f $$ (i,j) = (1-f(i div 2)) * U\<^sub>f $$ (i,j)" using assms(1,3) by simp moreover have f0: "U\<^sub>f $$ (i,i+1) * U\<^sub>f $$ (i+1,j) = f(i div 2) * U\<^sub>f $$ (i+1,j)" by (metis One_nat_def Suc_leI add.right_neutral add_Suc_right assms(1) assms(3) diff_add_inverse2 even_add even_mult_iff jozsa_transform_coeff(3) jozsa_transform_dim le_add2 le_eq_less_or_eq odd_one one_add_one power.simps(2)) ultimately have f1: "(U\<^sub>f * U\<^sub>f) $$ (i,j) = (1-f(i div 2)) * U\<^sub>f $$ (i,j) + f(i div 2) * U\<^sub>f $$ (i+1,j)" by auto thus ?thesis proof (induct rule: disj_four_cases) show "j=i \ (j=i+1 \ odd j) \ (j=i-1 \ even j \ i\1) \ (j\i \ \(j=i+1 \ odd j) \ \ (j=i-1 \ even j \ i\1))" by linarith next assume a0:"j=i" then have "U\<^sub>f $$ (i,j) = (1-f(i div 2))" using assms(1,2) a0 by simp moreover have "U\<^sub>f $$ (i+1,j) = f(i div 2)" using assms(1,3) a0 by auto ultimately have "(U\<^sub>f * U\<^sub>f) $$ (i,j) = (1-f(i div 2)) * (1-f(i div 2)) + f(i div 2) * f(i div 2)" using f1 by simp moreover have "(1-f(i div 2)) * (1-f(i div 2)) + f(i div 2) * f(i div 2) = 1" using f_values assms(1) by (metis (no_types, lifting) Nat.minus_nat.diff_0 diff_add_0 diff_add_inverse jozsa_transform_dim(1) less_power_add_imp_div_less mem_Collect_eq mult_eq_if one_power2 power2_eq_square power_one_right) ultimately show "(U\<^sub>f * U\<^sub>f) $$ (i,j) = 1\<^sub>m (dim_col U\<^sub>f) $$ (i,j)" by(metis assms(2) a0 index_one_mat(1) of_nat_1) next assume a0: "(j=i+1 \ odd j)" then have "U\<^sub>f $$ (i,j) = f(i div 2)" using assms(1,2) a0 by simp moreover have "U\<^sub>f $$ (i+1,j) = (1-f(i div 2))" using assms(2,3) a0 by simp ultimately have "(U\<^sub>f * U\<^sub>f) $$ (i,j) = (1-f(i div 2)) * f(i div 2) + f(i div 2) * (1-f(i div 2))" using f0 f1 assms by simp then show "(U\<^sub>f * U\<^sub>f) $$ (i,j) = 1\<^sub>m (dim_col U\<^sub>f) $$ (i,j)" using assms(1,2) a0 by auto next assume "(j=i-1 \ even j \ i\1)" then show "(U\<^sub>f * U\<^sub>f) $$ (i,j) = 1\<^sub>m (dim_col U\<^sub>f) $$ (i,j)" using assms(3) dvd_diffD1 odd_one by blast next assume a0:"(j\i \ \(j=i+1 \ odd j) \ \ (j=i-1 \ even j \ i\1))" then have "U\<^sub>f $$ (i,j) = 0" using assms(1,2) by(metis index_transpose_mat(1) jozsa_transform_coeff_is_zero jozsa_transform_dim transpose_of_jozsa_transform) moreover have "U\<^sub>f $$ (i+1,j) = 0" using assms a0 by auto ultimately have "(U\<^sub>f * U\<^sub>f) $$ (i,j) = (1-f(i div 2)) * 0 + f(i div 2) * 0" by (simp add: f1) then show "(U\<^sub>f * U\<^sub>f) $$ (i,j) = 1\<^sub>m (dim_col U\<^sub>f) $$ (i,j)" using a0 assms(1,2) by(metis add.left_neutral index_one_mat(1) jozsa_transform_dim mult_0_right of_nat_0) qed qed lemma (in jozsa) jozsa_transform_is_unitary_index_odd: fixes i j:: nat assumes "i < dim_row U\<^sub>f" and "j < dim_col U\<^sub>f" and "odd i" shows "(U\<^sub>f * U\<^sub>f) $$ (i,j) = 1\<^sub>m (dim_col U\<^sub>f) $$ (i,j)" proof- have f0: "i \ 1" using linorder_not_less assms(3) by auto have "(U\<^sub>f * U\<^sub>f) $$ (i,j) = (\k \ {i-1,i}. U\<^sub>f $$ (i,k) * U\<^sub>f $$ (k,j)) " using U\<^sub>f_mult_without_empty_summands_odd[of i j U\<^sub>f ] assms by simp moreover have "(\k \ {i-1,i}. U\<^sub>f $$ (i,k) * U\<^sub>f $$ (k,j)) = U\<^sub>f $$ (i,i-1) * U\<^sub>f $$ (i-1,j) + U\<^sub>f $$ (i,i) * U\<^sub>f $$ (i,j)" using f0 by simp moreover have "U\<^sub>f $$ (i,i) * U\<^sub>f $$ (i,j) = (1-f(i div 2)) * U\<^sub>f $$ (i,j)" using assms(1,2) by simp moreover have f1: "U\<^sub>f $$ (i,i-1) * U\<^sub>f $$ (i-1,j) = f(i div 2) * U\<^sub>f $$ (i-1,j)" using assms(1) assms(3) by simp ultimately have f2: "(U\<^sub>f * U\<^sub>f) $$ (i,j) = f(i div 2) * U\<^sub>f $$ (i-1,j) + (1-f(i div 2)) * U\<^sub>f $$ (i,j)" by simp then show "?thesis" proof (induct rule: disj_four_cases) show "j=i \ (j=i+1 \ odd j) \ (j=i-1 \ even j \ i\1) \ (j\i \ \(j=i+1 \ odd j) \ \ (j=i-1 \ even j \ i\1))" by linarith next assume a0:"j=i" then have "U\<^sub>f $$ (i,j) = (1-f(i div 2))" using assms(1,2) by simp moreover have "U\<^sub>f $$ (i-1,j) = f(i div 2)" using a0 assms by (metis index_transpose_mat(1) jozsa_transform_coeff(2) less_imp_diff_less odd_two_times_div_two_nat odd_two_times_div_two_succ transpose_of_jozsa_transform) ultimately have "(U\<^sub>f * U\<^sub>f) $$ (i,j) = f(i div 2) * f(i div 2) + (1-f(i div 2)) * (1-f(i div 2))" using f2 by simp moreover have "f(i div 2) * f(i div 2) + (1-f(i div 2)) * (1-f(i div 2)) = 1" using f_values assms(1) by (metis (no_types, lifting) Nat.minus_nat.diff_0 diff_add_0 diff_add_inverse jozsa_transform_dim(1) less_power_add_imp_div_less mem_Collect_eq mult_eq_if one_power2 power2_eq_square power_one_right) ultimately show "(U\<^sub>f * U\<^sub>f) $$ (i,j) = 1\<^sub>m (dim_col U\<^sub>f) $$ (i,j)" by(metis assms(2) a0 index_one_mat(1) of_nat_1) next assume a0:"(j=i+1 \ odd j)" then show "(U\<^sub>f * U\<^sub>f) $$ (i,j) = 1\<^sub>m (dim_col U\<^sub>f) $$ (i,j)" using assms(3) dvd_diffD1 odd_one even_plus_one_iff by blast next assume a0:"(j=i-1 \ even j \ i\1)" then have "(U\<^sub>f * U\<^sub>f) $$ (i,j) = f(i div 2) * (1-f(i div 2)) + (1-f(i div 2)) * f(i div 2)" using f0 f1 f2 assms by (metis jozsa_transform_coeff(1) Groups.ab_semigroup_mult_class.mult.commute even_succ_div_two f2 jozsa_transform_dim odd_two_times_div_two_nat odd_two_times_div_two_succ of_nat_add of_nat_mult) then show "(U\<^sub>f * U\<^sub>f) $$ (i,j) = 1\<^sub>m (dim_col U\<^sub>f) $$ (i,j)" using assms(1) a0 by auto next assume a0:"j\i \ \(j=i+1 \ odd j) \ \ (j=i-1 \ even j \ i\1)" then have "U\<^sub>f $$ (i,j) = 0" by (metis assms(1,2) index_transpose_mat(1) jozsa_transform_coeff_is_zero jozsa_transform_dim transpose_of_jozsa_transform) moreover have "U\<^sub>f $$ (i-1,j) = 0" using assms a0 f0 - by auto (smt One_nat_def Suc_n_not_le_n add_diff_inverse_nat assms(1) assms(2) diff_Suc_less even_add + by auto (smt (verit) One_nat_def Suc_n_not_le_n add_diff_inverse_nat assms(1) assms(2) diff_Suc_less even_add jozsa_transform_coeff_is_zero jozsa_axioms less_imp_le less_le_trans less_one odd_one) ultimately have "(U\<^sub>f * U\<^sub>f) $$ (i,j) = (1-f(i div 2)) * 0 + f(i div 2) * 0" using f2 by simp then show "(U\<^sub>f * U\<^sub>f) $$ (i,j) = 1\<^sub>m (dim_col U\<^sub>f) $$ (i,j)" using a0 assms by (metis add.left_neutral index_one_mat(1) jozsa_transform_dim mult_0_right of_nat_0) qed qed lemma (in jozsa) jozsa_transform_is_gate: shows "gate (n+1) U\<^sub>f" proof show "dim_row U\<^sub>f = 2^(n+1)" by simp next show "square_mat U\<^sub>f" by simp next show "unitary U\<^sub>f" proof- have "U\<^sub>f * U\<^sub>f = 1\<^sub>m (dim_col U\<^sub>f)" proof show "dim_row (U\<^sub>f * U\<^sub>f) = dim_row (1\<^sub>m (dim_col U\<^sub>f))" by simp show "dim_col (U\<^sub>f * U\<^sub>f) = dim_col (1\<^sub>m (dim_col U\<^sub>f))" by simp fix i j:: nat assume "i < dim_row (1\<^sub>m (dim_col U\<^sub>f))" and "j < dim_col (1\<^sub>m (dim_col U\<^sub>f))" then have "i < dim_row U\<^sub>f" and "j < dim_col U\<^sub>f" by auto then show "(U\<^sub>f * U\<^sub>f) $$ (i,j) = 1\<^sub>m (dim_col U\<^sub>f) $$ (i,j)" using jozsa_transform_is_unitary_index_odd jozsa_transform_is_unitary_index_even by blast qed thus ?thesis by (simp add: adjoint_of_jozsa_transform unitary_def) qed qed text \N-fold application of the tensor product\ fun iter_tensor:: "complex Matrix.mat \ nat \ complex Matrix.mat" ("_ \\<^bsup>_\<^esup>" 75) where "A \\<^bsup>(Suc 0)\<^esup> = A" | "A \\<^bsup>(Suc k)\<^esup> = A \ (A \\<^bsup>k\<^esup>)" lemma one_tensor_is_id [simp]: fixes A shows "A \\<^bsup>1\<^esup> = A" using one_mat_def by simp lemma iter_tensor_suc: fixes n assumes "n \ 1" shows " A \\<^bsup>(Suc n)\<^esup> = A \ (A \\<^bsup>n\<^esup>)" using assms by (metis Deutsch_Jozsa.iter_tensor.simps(2) One_nat_def Suc_le_D) lemma dim_row_of_iter_tensor [simp]: fixes A n assumes "n \ 1" shows "dim_row(A \\<^bsup>n\<^esup>) = (dim_row A)^n" using assms proof (rule nat_induct_at_least) show "dim_row (A \\<^bsup>1\<^esup>) = (dim_row A)^1" using one_tensor_is_id by simp fix n:: nat assume "n \ 1" and "dim_row (A \\<^bsup>n\<^esup>) = (dim_row A)^n" then show "dim_row (A \\<^bsup>Suc n\<^esup>) = (dim_row A)^Suc n" using iter_tensor_suc assms dim_row_tensor_mat by simp qed lemma dim_col_of_iter_tensor [simp]: fixes A n assumes "n \ 1" shows "dim_col(A \\<^bsup>n\<^esup>) = (dim_col A)^n" using assms proof (rule nat_induct_at_least) show "dim_col (A \\<^bsup>1\<^esup>) = (dim_col A)^1" using one_tensor_is_id by simp fix n:: nat assume "n \ 1" and "dim_col (A \\<^bsup>n\<^esup>) = (dim_col A)^n" then show "dim_col (A \\<^bsup>Suc n\<^esup>) = (dim_col A)^Suc n" using iter_tensor_suc assms dim_col_tensor_mat by simp qed lemma iter_tensor_values: fixes A n i j assumes "n \ 1" and "i < dim_row (A \ (A \\<^bsup>n\<^esup>))" and "j < dim_col (A \ (A \\<^bsup>n\<^esup>))" shows "(A \\<^bsup>(Suc n)\<^esup>) $$ (i,j) = (A \ (A \\<^bsup>n\<^esup>)) $$ (i,j)" using assms by (metis One_nat_def le_0_eq not0_implies_Suc iter_tensor.simps(2)) lemma iter_tensor_mult_distr: assumes "n \ 1" and "dim_col A = dim_row B" and "dim_col A > 0" and "dim_col B > 0" shows "(A \\<^bsup>(Suc n)\<^esup>) * (B \\<^bsup>(Suc n)\<^esup>) = (A * B) \ ((A \\<^bsup>n\<^esup>) * (B \\<^bsup>n\<^esup>))" proof- have "(A \\<^bsup>(Suc n)\<^esup>) * (B \\<^bsup>(Suc n)\<^esup>) = (A \ (A \\<^bsup>n\<^esup>)) * (B \ (B \\<^bsup>n\<^esup>))" using Suc_le_D assms(1) by fastforce then show "?thesis" using mult_distr_tensor[of "A" "B" "(iter_tensor A n)" "(iter_tensor B n)"] assms by simp qed lemma index_tensor_mat_with_vec2_row_cond: fixes A B:: "complex Matrix.mat" and i:: "nat" assumes "i < 2 * (dim_row B)" and "i \ dim_row B" and "dim_col B > 0" and "dim_row A = 2" and "dim_col A = 1" shows "(A \ B) $$ (i,0) = (A $$ (1,0)) * (B $$ (i-dim_row B,0))" proof- have "(A \ B) $$ (i,0) = A $$ (i div (dim_row B),0) * B $$ (i mod (dim_row B),0)" using assms index_tensor_mat[of A "dim_row A" "dim_col A" B "dim_row B" "dim_col B" i 0] by simp moreover have "i div (dim_row B) = 1" using assms(1,2,4) by simp then have "i mod (dim_row B) = i - (dim_row B)" by (simp add: modulo_nat_def) ultimately show "(A \ B) $$ (i,0) = (A $$ (1,0)) * (B $$ (i-dim_row B,0))" by (simp add: \i div dim_row B = 1\) qed lemma iter_tensor_of_gate_is_gate: fixes A:: "complex Matrix.mat" and n m:: "nat" assumes "gate m A" and "n \ 1" shows "gate (m*n) (A \\<^bsup>n\<^esup>)" using assms(2) proof(rule nat_induct_at_least) show "gate (m * 1) (A \\<^bsup>1\<^esup>)" using assms(1) by simp fix n:: nat assume "n \ 1" and IH:"gate (m * n) (A \\<^bsup>n\<^esup>)" then have "A \\<^bsup>(Suc n)\<^esup> = A \ (A \\<^bsup>n\<^esup>)" by (simp add: iter_tensor_suc) moreover have "gate (m*n + m) (A \\<^bsup>(Suc n)\<^esup>)" using tensor_gate assms(1) by (simp add: IH add.commute calculation(1)) then show "gate (m*(Suc n)) (A \\<^bsup>(Suc n)\<^esup>)" by (simp add: add.commute) qed lemma iter_tensor_of_state_is_state: fixes A:: "complex Matrix.mat" and n m:: "nat" assumes "state m A" and "n\1" shows "state (m*n) (A \\<^bsup>n\<^esup>)" using assms(2) proof(rule nat_induct_at_least) show "state (m * 1) (A \\<^bsup>1\<^esup>)" using one_tensor_is_id assms(1) by simp fix n:: nat assume "n \ 1" and IH:"state (m * n) (A \\<^bsup>n\<^esup>)" then have "A \\<^bsup>(Suc n)\<^esup> = A \ (A \\<^bsup>n\<^esup>)" by (simp add: iter_tensor_suc) moreover have "state (m*n + m) (A \\<^bsup>(Suc n)\<^esup>)" using tensor_gate assms(1) by (simp add: IH add.commute calculation) then show "state (m*(Suc n)) (A \\<^bsup>(Suc n)\<^esup>)" by (simp add: add.commute) qed text \ We prepare n+1 qubits. The first n qubits in the state $|0\rangle$, the last one in the state $|1\rangle$. \ abbreviation \\<^sub>1\<^sub>0:: "nat \ complex Matrix.mat" where "\\<^sub>1\<^sub>0 n \ Matrix.mat (2^n) 1 (\(i,j). 1/(sqrt 2)^n)" lemma \\<^sub>1\<^sub>0_values: fixes i j n assumes "i < dim_row (\\<^sub>1\<^sub>0 n)" and "j < dim_col (\\<^sub>1\<^sub>0 n)" shows "(\\<^sub>1\<^sub>0 n) $$ (i,j) = 1/(sqrt 2)^n" using assms case_prod_conv by simp text \$H^{\otimes n}$ is applied to $|0\rangle^{\otimes n}$.\ lemma H_on_ket_zero: shows "(H * |zero\) = \\<^sub>1\<^sub>0 1" proof fix i j:: nat assume "i < dim_row (\\<^sub>1\<^sub>0 1)" and "j < dim_col (\\<^sub>1\<^sub>0 1)" then have f1: "i \ {0,1} \ j = 0" by (simp add: less_2_cases) then show "(H * |zero\) $$ (i,j) = (\\<^sub>1\<^sub>0 1) $$ (i,j)" by (auto simp add: times_mat_def scalar_prod_def H_def ket_vec_def) next show "dim_row (H * |zero\) = dim_row (\\<^sub>1\<^sub>0 1)" by (simp add: H_def) show "dim_col (H * |zero\) = dim_col (\\<^sub>1\<^sub>0 1)" using H_def by (simp add: ket_vec_def) qed lemma \\<^sub>1\<^sub>0_tensor: assumes "n \ 1" shows "(\\<^sub>1\<^sub>0 1) \ (\\<^sub>1\<^sub>0 n) = (\\<^sub>1\<^sub>0 (Suc n))" proof have "dim_row (\\<^sub>1\<^sub>0 1) * dim_row (\\<^sub>1\<^sub>0 n) = 2^(Suc n)" by simp then show "dim_row ((\\<^sub>1\<^sub>0 1) \ (\\<^sub>1\<^sub>0 n)) = dim_row (\\<^sub>1\<^sub>0 (Suc n))" by simp have "dim_col (\\<^sub>1\<^sub>0 1) * dim_col (\\<^sub>1\<^sub>0 n) = 1" by simp then show "dim_col ((\\<^sub>1\<^sub>0 1) \ (\\<^sub>1\<^sub>0 n)) = dim_col (\\<^sub>1\<^sub>0 (Suc n))" by simp next fix i j:: nat assume a0: "i < dim_row (\\<^sub>1\<^sub>0 (Suc n))" and a1: "j < dim_col (\\<^sub>1\<^sub>0 (Suc n))" then have f0: "j = 0" and f1: "i < 2^(Suc n)" by auto then have f2:"(\\<^sub>1\<^sub>0 (Suc n)) $$ (i,j) = 1/(sqrt 2)^(Suc n)" using \\<^sub>1\<^sub>0_values[of "i" "(Suc n)" "j"] a0 a1 by simp show "((\\<^sub>1\<^sub>0 1) \ (\\<^sub>1\<^sub>0 n)) $$ (i,j) = (\\<^sub>1\<^sub>0 (Suc n)) $$ (i,j)" proof (rule disjE) (*case distinction*) show "i < dim_row (\\<^sub>1\<^sub>0 n) \ i \ dim_row (\\<^sub>1\<^sub>0 n)" by linarith next (* case i < dim_row (\\<^sub>1\<^sub>0 n) *) assume a2: "i < dim_row (\\<^sub>1\<^sub>0 n)" then have "((\\<^sub>1\<^sub>0 1) \ (\\<^sub>1\<^sub>0 n)) $$ (i,j) = (\\<^sub>1\<^sub>0 1) $$ (0,0) * (\\<^sub>1\<^sub>0 n) $$ (i,0)" using index_tensor_mat f0 assms by simp also have "... = 1/sqrt(2) * 1/(sqrt(2)^n)" using \\<^sub>1\<^sub>0_values a2 assms by simp finally show "((\\<^sub>1\<^sub>0 1) \ (\\<^sub>1\<^sub>0 n)) $$ (i,j) = (\\<^sub>1\<^sub>0 (Suc n)) $$ (i,j)" using f2 divide_divide_eq_left power_Suc by simp next (* case i \ dim_row (\\<^sub>1\<^sub>0 n) *) assume "i \ dim_row (\\<^sub>1\<^sub>0 n)" then have "((\\<^sub>1\<^sub>0 1) \ (\\<^sub>1\<^sub>0 n)) $$ (i,0) = ((\\<^sub>1\<^sub>0 1) $$ (1, 0)) * ((\\<^sub>1\<^sub>0 n) $$ ( i -dim_row (\\<^sub>1\<^sub>0 n),0))" using index_tensor_mat_with_vec2_row_cond[of i "(\\<^sub>1\<^sub>0 1)" "(\\<^sub>1\<^sub>0 n)" ] a0 a1 f0 by (metis dim_col_mat(1) dim_row_mat(1) index_tensor_mat_with_vec2_row_cond power_Suc power_one_right) then have "((\\<^sub>1\<^sub>0 1) \ (\\<^sub>1\<^sub>0 n)) $$ (i,0) = 1/sqrt(2) * 1/(sqrt 2)^n" using \\<^sub>1\<^sub>0_values[of "i -dim_row (\\<^sub>1\<^sub>0 n)" "n" "j"] a0 a1 by simp then show "((\\<^sub>1\<^sub>0 1) \ (\\<^sub>1\<^sub>0 n)) $$ (i,j) = (\\<^sub>1\<^sub>0 (Suc n)) $$ (i,j)" using f0 f1 divide_divide_eq_left power_Suc by simp qed qed lemma \\<^sub>1\<^sub>0_tensor_is_state: assumes "n \ 1" shows "state n ( |zero\ \\<^bsup>n\<^esup>)" using iter_tensor_of_state_is_state ket_zero_is_state assms by fastforce lemma iter_tensor_of_H_is_gate: assumes "n \ 1" shows "gate n (H \\<^bsup>n\<^esup>)" using iter_tensor_of_gate_is_gate H_is_gate assms by fastforce lemma iter_tensor_of_H_on_zero_tensor: assumes "n \ 1" shows "(H \\<^bsup>n\<^esup>) * ( |zero\ \\<^bsup>n\<^esup>) = \\<^sub>1\<^sub>0 n" using assms proof(rule nat_induct_at_least) show "(H \\<^bsup>1\<^esup>) * ( |zero\ \\<^bsup>1\<^esup>) = \\<^sub>1\<^sub>0 1" using H_on_ket_zero by simp next fix n:: nat assume a0: "n \ 1" and IH: "(H \\<^bsup>n\<^esup>) * ( |zero\ \\<^bsup>n\<^esup>) = \\<^sub>1\<^sub>0 n" then have "(H \\<^bsup>(Suc n)\<^esup>) * ( |zero\ \\<^bsup>(Suc n)\<^esup>) = (H * |zero\) \ ((H \\<^bsup>n\<^esup>) * ( |zero\ \\<^bsup>n\<^esup>))" using iter_tensor_mult_distr[of "n" "H" "|zero\"] a0 ket_vec_def H_def by(simp add: H_def) also have "... = (H * |zero\) \ (\\<^sub>1\<^sub>0 n)" using IH by simp also have "... = (\\<^sub>1\<^sub>0 1) \ (\\<^sub>1\<^sub>0 n)" using H_on_ket_zero by simp also have "... = (\\<^sub>1\<^sub>0 (Suc n))" using \\<^sub>1\<^sub>0_tensor a0 by simp finally show "(H \\<^bsup>(Suc n)\<^esup>) * ( |zero\ \\<^bsup>(Suc n)\<^esup>) = (\\<^sub>1\<^sub>0 (Suc n))" by simp qed lemma \\<^sub>1\<^sub>0_is_state: assumes "n \ 1" shows "state n (\\<^sub>1\<^sub>0 n)" using iter_tensor_of_H_is_gate \\<^sub>1\<^sub>0_tensor_is_state assms gate_on_state_is_state iter_tensor_of_H_on_zero_tensor assms by metis abbreviation \\<^sub>1\<^sub>1:: "complex Matrix.mat" where "\\<^sub>1\<^sub>1 \ Matrix.mat 2 1 (\(i,j). if i=0 then 1/sqrt(2) else -1/sqrt(2))" lemma H_on_ket_one_is_\\<^sub>1\<^sub>1: shows "(H * |one\) = \\<^sub>1\<^sub>1" proof fix i j:: nat assume "i < dim_row \\<^sub>1\<^sub>1" and "j < dim_col \\<^sub>1\<^sub>1" then have "i \ {0,1} \ j = 0" by (simp add: less_2_cases) then show "(H * |one\) $$ (i,j) = \\<^sub>1\<^sub>1 $$ (i,j)" by (auto simp add: times_mat_def scalar_prod_def H_def ket_vec_def) next show "dim_row (H * |one\) = dim_row \\<^sub>1\<^sub>1" by (simp add: H_def) next show "dim_col (H * |one\) = dim_col \\<^sub>1\<^sub>1" by (simp add: H_def ket_vec_def) qed abbreviation \\<^sub>1:: "nat \ complex Matrix.mat" where "\\<^sub>1 n \ Matrix.mat (2^(n+1)) 1 (\(i,j). if even i then 1/(sqrt 2)^(n+1) else -1/(sqrt 2)^(n+1))" lemma \\<^sub>1_values_even[simp]: fixes i j n assumes "i < dim_row (\\<^sub>1 n)" and "j < dim_col (\\<^sub>1 n)" and "even i" shows "(\\<^sub>1 n) $$ (i,j) = 1/(sqrt 2)^(n+1)" using assms case_prod_conv by simp lemma \\<^sub>1_values_odd [simp]: fixes i j n assumes "i < dim_row (\\<^sub>1 n)" and "j < dim_col (\\<^sub>1 n)" and "odd i" shows "(\\<^sub>1 n) $$ (i,j) = -1/(sqrt 2)^(n+1)" using assms case_prod_conv by simp lemma "\\<^sub>1\<^sub>0_tensor_\\<^sub>1\<^sub>1_is_\\<^sub>1": assumes "n \ 1" shows "(\\<^sub>1\<^sub>0 n) \ \\<^sub>1\<^sub>1 = \\<^sub>1 n" proof show "dim_col ((\\<^sub>1\<^sub>0 n) \ \\<^sub>1\<^sub>1) = dim_col (\\<^sub>1 n)" by simp next show "dim_row ((\\<^sub>1\<^sub>0 n) \ \\<^sub>1\<^sub>1) = dim_row (\\<^sub>1 n)" by simp next fix i j:: nat assume a0: "i < dim_row (\\<^sub>1 n)" and a1: "j < dim_col (\\<^sub>1 n)" then have "i < 2^(n+1)" and "j = 0" by auto then have f0: "((\\<^sub>1\<^sub>0 n) \ \\<^sub>1\<^sub>1) $$ (i,j) = 1/(sqrt 2)^n * \\<^sub>1\<^sub>1 $$ (i mod 2, j)" using \\<^sub>1\<^sub>0_values[of "i div 2" n "j div 1"] a0 a1 by simp show "((\\<^sub>1\<^sub>0 n) \ \\<^sub>1\<^sub>1) $$ (i,j) = (\\<^sub>1 n) $$ (i,j)" using f0 \\<^sub>1_values_even \\<^sub>1_values_odd a0 a1 by auto qed lemma \\<^sub>1_is_state: assumes "n \ 1" shows "state (n+1) (\\<^sub>1 n)" using assms \\<^sub>1\<^sub>0_tensor_\\<^sub>1\<^sub>1_is_\\<^sub>1 \\<^sub>1\<^sub>0_is_state H_on_ket_one_is_state H_on_ket_one_is_\\<^sub>1\<^sub>1 tensor_state by metis abbreviation (in jozsa) \\<^sub>2:: "complex Matrix.mat" where "\\<^sub>2 \ Matrix.mat (2^(n+1)) 1 (\(i,j). if even i then (-1)^f(i div 2)/(sqrt 2)^(n+1) else (-1)^(f(i div 2)+1)/(sqrt 2)^(n+1))" lemma (in jozsa) \\<^sub>2_values_even [simp]: fixes i j assumes "i < dim_row \\<^sub>2 " and "j < dim_col \\<^sub>2" and "even i" shows "\\<^sub>2 $$ (i,j) = (-1)^f(i div 2)/(sqrt 2)^(n+1)" using assms case_prod_conv by simp lemma (in jozsa) \\<^sub>2_values_odd [simp]: fixes i j assumes "i < dim_row \\<^sub>2" and "j < dim_col \\<^sub>2" and "odd i" shows "\\<^sub>2 $$ (i,j) = (-1)^(f(i div 2)+1)/(sqrt 2)^(n+1)" using assms case_prod_conv by simp lemma (in jozsa) \\<^sub>2_values_odd_hidden [simp]: assumes "2*k+1 < dim_row \\<^sub>2" and "j < dim_col \\<^sub>2" shows "\\<^sub>2 $$ (2*k+1,j) = ((-1)^(f((2*k+1) div 2)+1))/(sqrt 2)^(n+1)" using assms by simp lemma (in jozsa) snd_rep_of_\\<^sub>2: assumes "i < dim_row \\<^sub>2" shows "((1-f(i div 2)) + -f(i div 2)) * 1/(sqrt 2)^(n+1) = (-1)^f(i div 2)/(sqrt 2)^(n+1)" and "(-(1-f(i div 2))+(f(i div 2)))* 1/(sqrt 2)^(n+1) = (-1)^(f(i div 2)+1)/(sqrt 2)^(n+1)" proof- have "i div 2 \ {i. i < 2 ^ n}" using assms by auto then have "real (Suc 0 - f (i div 2)) - real (f (i div 2)) = (- 1) ^ f (i div 2)" using assms f_values by auto thus "((1-f(i div 2)) + -f(i div 2)) * 1/(sqrt 2)^(n+1) = (-1)^f(i div 2)/(sqrt 2)^(n+1)" by auto next have "i div 2 \ {i. i < 2^n}" using assms by simp then have "(real (f (i div 2)) - real (Suc 0 - f (i div 2))) / (sqrt 2 ^ (n+1)) = - ((- 1) ^ f (i div 2) / (sqrt 2 ^ (n+1)))" using assms f_values by fastforce then show "(-(1-f(i div 2))+(f(i div 2)))* 1/(sqrt 2)^(n+1) = (-1)^(f(i div 2)+1)/(sqrt 2)^(n+1)" by simp qed lemma (in jozsa) jozsa_transform_times_\\<^sub>1_is_\\<^sub>2: shows "U\<^sub>f * (\\<^sub>1 n) = \\<^sub>2" proof show "dim_row (U\<^sub>f * (\\<^sub>1 n)) = dim_row \\<^sub>2" by simp next show "dim_col (U\<^sub>f * (\\<^sub>1 n)) = dim_col \\<^sub>2" by simp next fix i j ::nat assume a0: "i < dim_row \\<^sub>2" and a1: "j < dim_col \\<^sub>2" then have f0:"i \ {0..2^(n+1)} \ j=0" by simp then have f1: "i < dim_row U\<^sub>f \ j < dim_col U\<^sub>f " using a0 by simp have f2: "i < dim_row (\\<^sub>1 n) \ j < dim_col (\\<^sub>1 n)" using a0 a1 by simp show "(U\<^sub>f * (\\<^sub>1 n)) $$ (i,j) = \\<^sub>2 $$ (i,j)" proof (rule disjE) show "even i \ odd i" by auto next assume a2: "even i" then have "(U\<^sub>f * (\\<^sub>1 n)) $$ (i,j) = (\k \ {i,i+1}. U\<^sub>f $$ (i,k) * (\\<^sub>1 n) $$ (k,j))" using f1 f2 U\<^sub>f_mult_without_empty_summands_even[of i j "(\\<^sub>1 n)"] by simp moreover have "U\<^sub>f $$ (i,i) * (\\<^sub>1 n) $$ (i,j) = (1-f(i div 2))* 1/(sqrt 2)^(n+1)" using f0 f1 a2 by simp moreover have "U\<^sub>f $$ (i,i+1) * (\\<^sub>1 n) $$ (i+1,j) = (-f(i div 2))* 1/(sqrt 2)^(n+1)" using f0 f1 a2 by auto ultimately have "(U\<^sub>f * (\\<^sub>1 n)) $$ (i,j) = (1-f(i div 2))* 1/(sqrt 2)^(n+1) + (-f(i div 2))* 1/(sqrt 2)^(n+1)" by simp also have "... = ((1-f(i div 2))+-f(i div 2)) * 1/(sqrt 2)^(n+1)" using add_divide_distrib by (metis (no_types, opaque_lifting) mult.right_neutral of_int_add of_int_of_nat_eq) also have "... = \\<^sub>2 $$ (i,j)" using a0 a1 a2 snd_rep_of_\\<^sub>2 by simp finally show "(U\<^sub>f * (\\<^sub>1 n)) $$ (i,j) = \\<^sub>2 $$ (i,j)" by simp next assume a2: "odd i" then have f6: "i\1" using linorder_not_less by auto have "(U\<^sub>f * (\\<^sub>1 n)) $$ (i,j) = (\k \ {i-1,i}. U\<^sub>f $$ (i,k) * (\\<^sub>1 n) $$ (k,j))" using f1 f2 a2 U\<^sub>f_mult_without_empty_summands_odd[of i j "(\\<^sub>1 n)"] by (metis dim_row_mat(1) jozsa_transform_dim(2)) moreover have "(\k \ {i-1,i}. U\<^sub>f $$ (i,k) * (\\<^sub>1 n) $$ (k,j)) = U\<^sub>f $$ (i,i-1) * (\\<^sub>1 n) $$ (i-1,j) + U\<^sub>f $$ (i,i) * (\\<^sub>1 n) $$ (i,j)" using a2 f6 by simp moreover have "U\<^sub>f $$ (i,i) * (\\<^sub>1 n) $$ (i,j) = (1-f(i div 2))* -1/(sqrt 2)^(n+1)" using f1 f2 a2 by simp moreover have "U\<^sub>f $$ (i,i-1) * (\\<^sub>1 n) $$ (i-1,j) = f(i div 2)* 1/(sqrt 2)^(n+1)" using a0 a1 a2 by simp ultimately have "(U\<^sub>f * (\\<^sub>1 n)) $$ (i,j) = (1-f(i div 2))* -1/(sqrt 2)^(n+1) +(f(i div 2))* 1/(sqrt 2)^(n+1)" using of_real_add by simp also have "... = (-(1-f(i div 2)) + (f(i div 2))) * 1/(sqrt 2)^(n+1)" by (metis (no_types, opaque_lifting) mult.right_neutral add_divide_distrib mult_minus1_right of_int_add of_int_of_nat_eq) also have "... = (-1)^(f(i div 2)+1)/(sqrt 2)^(n+1)" using a0 a1 a2 snd_rep_of_\\<^sub>2 by simp finally show "(U\<^sub>f * (\\<^sub>1 n)) $$ (i,j) = \\<^sub>2 $$ (i,j)" using a0 a1 a2 by simp qed qed lemma (in jozsa) \\<^sub>2_is_state: shows "state (n+1) \\<^sub>2" using jozsa_transform_times_\\<^sub>1_is_\\<^sub>2 jozsa_transform_is_gate \\<^sub>1_is_state dim gate_on_state_is_state by fastforce text \@{text "H^\<^sub>\ n"} is the result of taking the nth tensor product of H\ abbreviation iter_tensor_of_H_rep:: "nat \ complex Matrix.mat" ("H^\<^sub>\ _") where "iter_tensor_of_H_rep n \ Matrix.mat (2^n) (2^n) (\(i,j).(-1)^(i \\<^bsub>n\<^esub> j)/(sqrt 2)^n)" lemma tensor_of_H_values [simp]: fixes n i j:: nat assumes "i < dim_row (H^\<^sub>\ n)" and "j < dim_col (H^\<^sub>\ n)" shows "(H^\<^sub>\ n) $$ (i,j) = (-1)^(i \\<^bsub>n\<^esub> j)/(sqrt 2)^n" using assms by simp lemma dim_row_of_iter_tensor_of_H [simp]: assumes "n \ 1" shows "1 < dim_row (H^\<^sub>\ n)" using assms by(metis One_nat_def Suc_1 dim_row_mat(1) le_trans lessI linorder_not_less one_less_power) lemma iter_tensor_of_H_fst_pos: fixes n i j:: nat assumes "i < 2^n \ j < 2^n" and "i < 2^(n+1) \ j < 2^(n+1)" shows "(H^\<^sub>\ (Suc n)) $$ (i,j) = 1/sqrt(2) * ((H^\<^sub>\ n) $$ (i mod 2^n, j mod 2^n))" proof- have "(H^\<^sub>\ (Suc n)) $$ (i,j) = (-1)^(bip i (Suc n) j)/(sqrt 2)^(Suc n)" using assms by simp moreover have "bip i (Suc n) j = bip (i mod 2^n) n (j mod 2^n)" using bitwise_inner_prod_fst_el_0 assms(1) by simp ultimately show ?thesis using bitwise_inner_prod_def by simp qed lemma iter_tensor_of_H_fst_neg: fixes n i j:: nat assumes "i \ 2^n \ j \ 2^n" and "i < 2^(n+1) \ j < 2^(n+1)" shows "(H^\<^sub>\ (Suc n)) $$ (i,j) = -1/sqrt(2) * (H^\<^sub>\ n) $$ (i mod 2^n, j mod 2^n)" proof- have "(H^\<^sub>\ (Suc n)) $$ (i,j) = (-1)^(bip i (n+1) j)/(sqrt 2)^(n+1)" using assms(2) by simp moreover have "bip i (n+1) j = 1 + bip (i mod 2^n) n (j mod 2^n)" using bitwise_inner_prod_fst_el_is_1 assms by simp ultimately show ?thesis by simp qed lemma H_tensor_iter_tensor_of_H: fixes n:: nat shows "(H \ H^\<^sub>\ n) = H^\<^sub>\ (Suc n)" proof fix i j:: nat assume a0: "i < dim_row (H^\<^sub>\ (Suc n))" and a1: "j < dim_col (H^\<^sub>\ (Suc n))" then have f0: "i \ {0..<2^(n+1)} \ j \ {0..<2^(n+1)}" by simp then have f1: "(H \ H^\<^sub>\ n) $$ (i,j) = H $$ (i div (dim_row (H^\<^sub>\ n)),j div (dim_col (H^\<^sub>\ n))) * (H^\<^sub>\ n) $$ (i mod (dim_row (H^\<^sub>\ n)),j mod (dim_col (H^\<^sub>\ n)))" by (simp add: H_without_scalar_prod) show "(H \ H^\<^sub>\ n) $$ (i,j) = (H^\<^sub>\ (Suc n)) $$ (i,j)" proof (rule disjE) show "(i < 2^n \ j < 2^n) \ \(i < 2^n \ j < 2^n)" by auto next assume a2: "(i < 2^n \ j < 2^n)" then have "(H^\<^sub>\ (Suc n)) $$ (i,j) = 1/sqrt(2) * ((H^\<^sub>\ n) $$ (i mod 2^n, j mod 2^n))" using a0 a1 f0 iter_tensor_of_H_fst_pos by (metis (mono_tags, lifting) atLeastLessThan_iff) moreover have "H $$ (i div (dim_row (H^\<^sub>\ n)),j div (dim_col (H^\<^sub>\ n))) = 1/sqrt 2" using a0 a1 f0 H_without_scalar_prod H_values a2 by (metis (no_types, lifting) dim_col_mat(1) dim_row_mat(1) div_less le_eq_less_or_eq le_numeral_extra(2) less_power_add_imp_div_less plus_1_eq_Suc power_one_right) ultimately show "(H \ H^\<^sub>\ n) $$ (i,j) = (H^\<^sub>\ (Suc n)) $$ (i,j)" using f1 by simp next assume a2: "\(i < 2^n \ j < 2^n)" then have "i \ 2^n \ j \ 2^n" by simp then have f2:"(H^\<^sub>\ (Suc n)) $$ (i,j) = -1/sqrt(2) * ((H^\<^sub>\ n) $$ (i mod 2^n, j mod 2^n))" using a0 a1 f0 iter_tensor_of_H_fst_neg by simp have "i div (dim_row (H^\<^sub>\ n)) =1" and "j div (dim_row (H^\<^sub>\ n)) = 1" using a2 a0 a1 by auto then have "H $$ (i div (dim_row (H^\<^sub>\ n)),j div (dim_col (H^\<^sub>\ n))) = -1/sqrt 2" using a0 a1 f0 H_values_right_bottom[of "i div (dim_row (H^\<^sub>\ n))" "j div (dim_col (H^\<^sub>\ n))"] a2 by fastforce then show "(H \ H^\<^sub>\ n) $$ (i,j) = (H^\<^sub>\ (Suc n)) $$ (i,j)" using f1 f2 by simp qed next show "dim_row (H \ H^\<^sub>\ n) = dim_row (H^\<^sub>\ (Suc n))" by (simp add: H_without_scalar_prod) next show "dim_col (H \ H^\<^sub>\ n) = dim_col (H^\<^sub>\ (Suc n))" by (simp add: H_without_scalar_prod) qed text \ We prove that @{term "H^\<^sub>\ n"} is indeed the matrix representation of @{term "H \\<^bsup>n\<^esup>"}, the iterated tensor product of the Hadamard gate H. \ lemma one_tensor_of_H_is_H: shows "(H^\<^sub>\ 1) = H" proof(rule eq_matI) show "dim_row (H^\<^sub>\ 1) = dim_row H" by (simp add: H_without_scalar_prod) show "dim_col (H^\<^sub>\ 1) = dim_col H" by (simp add: H_without_scalar_prod) next fix i j:: nat assume a0:"i < dim_row H" and a1:"j < dim_col H" then show "(H^\<^sub>\ 1) $$ (i,j) = H $$ (i,j)" proof- have "(H^\<^sub>\ 1) $$ (0, 0) = 1/sqrt(2)" using bitwise_inner_prod_def bin_rep_def by simp moreover have "(H^\<^sub>\ 1) $$ (0,1) = 1/sqrt(2)" using bitwise_inner_prod_def bin_rep_def by simp moreover have "(H^\<^sub>\ 1) $$ (1,0) = 1/sqrt(2)" using bitwise_inner_prod_def bin_rep_def by simp moreover have "(H^\<^sub>\ 1) $$ (1,1) = -1/sqrt(2)" using bitwise_inner_prod_def bin_rep_def by simp ultimately show "(H^\<^sub>\ 1) $$ (i,j) = H $$ (i,j)" using a0 a1 H_values H_values_right_bottom by (metis (no_types, lifting) H_without_scalar_prod One_nat_def dim_col_mat(1) dim_row_mat(1) divide_minus_left less_2_cases) qed qed lemma iter_tensor_of_H_rep_is_correct: fixes n:: nat assumes "n \ 1" shows "(H \\<^bsup>n\<^esup>) = H^\<^sub>\ n" using assms proof(rule nat_induct_at_least) show "(H \\<^bsup>1\<^esup>) = H^\<^sub>\ 1" using one_tensor_is_id one_tensor_of_H_is_H by simp next fix n:: nat assume a0:"n \ 1" and IH:"(H \\<^bsup>n\<^esup>) = H^\<^sub>\ n" then have "(H \\<^bsup>(Suc n)\<^esup>) = H \ (H \\<^bsup>n\<^esup>)" using iter_tensor_suc Nat.Suc_eq_plus1 by metis also have "... = H \ (H^\<^sub>\ n)" using IH by simp also have "... = H^\<^sub>\ (Suc n)" using a0 H_tensor_iter_tensor_of_H by simp finally show "(H \\<^bsup>(Suc n)\<^esup>) = H^\<^sub>\ (Suc n)" by simp qed text \@{text "HId^\<^sub>\ 1"} is the result of taking the tensor product of the nth tensor of H and Id 1 \ abbreviation tensor_of_H_tensor_Id:: "nat \ complex Matrix.mat" ("HId^\<^sub>\ _") where "tensor_of_H_tensor_Id n \ Matrix.mat (2^(n+1)) (2^(n+1)) (\(i,j). if (i mod 2 = j mod 2) then (-1)^((i div 2) \\<^bsub>n\<^esub> (j div 2))/(sqrt 2)^n else 0)" lemma mod_2_is_both_even_or_odd: "((even i \ even j) \ (odd i \ odd j)) \ (i mod 2 = j mod 2)" by (metis even_iff_mod_2_eq_zero odd_iff_mod_2_eq_one) lemma HId_values [simp]: assumes "n \ 1" and "i < dim_row (HId^\<^sub>\ n)" and "j < dim_col (HId^\<^sub>\ n)" shows "even i \ even j \ (HId^\<^sub>\ n) $$ (i,j) = (-1)^((i div 2) \\<^bsub>n\<^esub> (j div 2))/(sqrt 2)^n" and "odd i \ odd j \ (HId^\<^sub>\ n) $$ (i,j) = (-1)^((i div 2) \\<^bsub>n\<^esub> (j div 2))/(sqrt 2)^n" and "(i mod 2 = j mod 2) \ (HId^\<^sub>\ n) $$ (i,j) = (-1)^((i div 2) \\<^bsub>n\<^esub> (j div 2))/(sqrt 2)^n" and "\(i mod 2 = j mod 2) \ (HId^\<^sub>\ n) $$ (i,j) = 0" using assms mod_2_is_both_even_or_odd by auto lemma iter_tensor_of_H_tensor_Id_is_HId: shows "(H^\<^sub>\ n) \ Id 1 = HId^\<^sub>\ n" proof show "dim_row ((H^\<^sub>\ n) \ Id 1) = dim_row (HId^\<^sub>\ n)" by (simp add: Quantum.Id_def) show "dim_col ((H^\<^sub>\ n) \ Id 1) = dim_col (HId^\<^sub>\ n)" by (simp add: Quantum.Id_def) next fix i j:: nat assume a0: "i < dim_row (HId^\<^sub>\ n)" and a1: "j < dim_col (HId^\<^sub>\ n)" then have f0: "i < (2^(n+1)) \ j < (2^(n+1))" by simp then have "i < dim_row (H^\<^sub>\ n) * dim_row (Id 1) \ j < dim_col (H^\<^sub>\ n) * dim_col (Id 1)" using Id_def by simp moreover have "dim_col (H^\<^sub>\ n) \ 0 \ dim_col (Id 1) \ 0" using Id_def by simp ultimately have f1: "((H^\<^sub>\ n) \ (Id 1)) $$ (i,j) = (H^\<^sub>\ n) $$ (i div (dim_row (Id 1)),j div (dim_col (Id 1))) * (Id 1) $$ (i mod (dim_row (Id 1)),j mod (dim_col (Id 1)))" by (simp add: Quantum.Id_def) show "((H^\<^sub>\ n)\Id 1) $$ (i,j) = (HId^\<^sub>\ n) $$ (i,j)" proof (rule disjE) show "(i mod 2 = j mod 2) \ \ (i mod 2 = j mod 2)" by simp next assume a2:"(i mod 2 = j mod 2)" then have "(Id 1) $$ (i mod (dim_row (Id 1)),j mod (dim_col (Id 1))) = 1" by (simp add: Quantum.Id_def) moreover have "(H^\<^sub>\ n) $$ (i div (dim_row (Id 1)), j div (dim_col (Id 1))) = (-1)^((i div (dim_row (Id 1))) \\<^bsub>n\<^esub> (j div (dim_col (Id 1))))/(sqrt 2)^n" using tensor_of_H_values Id_def f0 less_mult_imp_div_less by simp ultimately show "((H^\<^sub>\ n) \ Id 1) $$ (i,j) = (HId^\<^sub>\ n) $$ (i,j)" using a2 f0 f1 Id_def by simp next assume a2: "\(i mod 2 = j mod 2)" then have "(Id 1) $$ (i mod (dim_row (Id 1)),j mod (dim_col (Id 1))) = 0" by (simp add: Quantum.Id_def) then show "((H^\<^sub>\ n) \ Id 1) $$ (i,j) = (HId^\<^sub>\ n) $$ (i,j)" using a2 f0 f1 by simp qed qed lemma HId_is_gate: assumes "n \ 1" shows "gate (n+1) (HId^\<^sub>\ n)" proof- have "(HId^\<^sub>\ n) = (H^\<^sub>\ n) \ Id 1" using iter_tensor_of_H_tensor_Id_is_HId by simp moreover have "gate 1 (Id 1)" using id_is_gate by simp moreover have "gate n (H^\<^sub>\ n)" using H_is_gate iter_tensor_of_gate_is_gate[of 1 H n] assms by(simp add: iter_tensor_of_H_rep_is_correct) ultimately show "gate (n+1) (HId^\<^sub>\ n)" using tensor_gate by presburger qed text \State @{term "\\<^sub>3"} is obtained by the multiplication of @{term "HId^\<^sub>\ n"} and @{term "\\<^sub>2"}\ abbreviation (in jozsa) \\<^sub>3:: "complex Matrix.mat" where "\\<^sub>3 \ Matrix.mat (2^(n+1)) 1 (\(i,j). if even i then (\k<2^n. (-1)^(f(k) + ((i div 2) \\<^bsub>n\<^esub> k))/((sqrt 2)^n * (sqrt 2)^(n+1))) else (\k<2^n. (-1)^(f(k)+ 1 + ((i div 2) \\<^bsub>n\<^esub> k)) /((sqrt 2)^n * (sqrt 2)^(n+1))))" lemma (in jozsa) \\<^sub>3_values: assumes "i < dim_row \\<^sub>3" shows "odd i \ \\<^sub>3 $$ (i,0) = (\k<2^n. (-1)^(f(k) + 1 + ((i div 2) \\<^bsub>n\<^esub> k))/((sqrt 2)^n * (sqrt 2)^(n+1)))" using assms by simp lemma (in jozsa) \\<^sub>3_dim [simp]: shows "1 < dim_row \\<^sub>3" using dim_row_mat(1) nat_neq_iff by fastforce lemma sum_every_odd_summand_is_zero: fixes n:: nat assumes "n \ 1" shows "\f::(nat \ complex).(\i. i<2^(n+1) \ odd i \ f i = 0) \ (\k\{0..<2^(n+1)}. f k) = (\k\{0..<2^n}. f (2*k))" using assms proof(rule nat_induct_at_least) show "\f::(nat \ complex).(\i. i<2^(1+1) \ odd i \ f i = 0) \ (\k\{0..<2^(1+1)}. f k) = (\k \ {0..<2^1}. f (2*k))" proof(rule allI,rule impI) fix f:: "(nat \ complex)" assume asm: "(\i. i<2^(1+1) \ odd i \ f i = 0)" moreover have "(\k\{0..<4}. f k) = f 0 + f 1 + f 2 + f 3" by (simp add: add.commute add.left_commute) moreover have "f 1 = 0" using asm by simp moreover have "f 3 = 0" using asm by simp moreover have "(\k\{0..<2^1}. f (2*k)) = f 0 + f 2" using add.commute add.left_commute by simp ultimately show "(\k\{0..<2^(1+1)}. f k) = (\k\{0..<2^1}. f (2*k))" by simp qed next fix n:: nat assume "n \ 1" and IH: "\f::(nat \complex).(\i. i<2^(n+1) \ odd i \ f i = 0) \ (\k\{0..<2^(n+1)}. f k) = (\k\{0..<2^n}. f (2*k))" show "\f::(nat \complex).(\i. i<2^(Suc n +1) \ odd i \ f i = 0) \ (\k\{0..<2^(Suc n +1)}. f k) = (\k\{0..< 2^(Suc n)}. f (2*k))" proof (rule allI,rule impI) fix f::"nat \ complex" assume asm: "(\i. i<2^(Suc n +1) \ odd i \ f i = 0)" have f0: "(\k\{0..<2^(n+1)}. f k) = (\k\{0..<2^n}. f (2*k))" using asm IH by simp have f1: "(\k\{0..<2^(n+1)}. (\x. f (x+2^(n+1))) k) = (\k\{0..< 2^n}. (\x. f (x+2^(n+1))) (2*k))" using asm IH by simp have "(\k\{0..<2^(n+2)}. f k) = (\k\{0..<2^(n+1)}. f k) + (\k\{2^(n+1)..<2^(n+2)}. f k)" by (simp add: sum.atLeastLessThan_concat) also have "... = (\k\{0..<2^n}. f (2*k)) + (\k\{2^(n+1)..<2^(n+2)}. f k)" using f0 by simp also have "... = (\k\{0..<2^n}. f (2*k)) + (\k\{0..<2^(n+1)}. f (k+2^(n+1)))" using sum.shift_bounds_nat_ivl[of "f" "0" "2^(n+1)" "2^(n+1)"] by simp also have "... = (\k\{0..<2^n}. f (2*k)) + (\k\{0..< 2^n}. (\x. f (x+2^(n+1))) (2*k))" using f1 by simp also have "... = (\k\{0..<2^n}. f (2*k)) + (\k\{2^n..< 2^(n+1)}. f (2 *k))" using sum.shift_bounds_nat_ivl[of "\x. (f::nat\complex) (2*(x-2^n)+2^(n+1))" "0" "2^n" "2^n"] by (simp add: mult_2) also have "... = (\k \ {0..<2^(n+1)}. f (2*k))" by (metis Suc_eq_plus1 lessI less_imp_le_nat one_le_numeral power_increasing sum.atLeastLessThan_concat zero_le) finally show "(\k\{0..<2^((Suc n)+1)}. f k) = (\k\{0..< 2^(Suc n)}. f (2*k))" by (metis Suc_eq_plus1 add_2_eq_Suc') qed qed lemma sum_every_even_summand_is_zero: fixes n:: nat assumes "n \ 1" shows "\f::(nat \ complex).(\i. i<2^(n+1) \ even i \ f i = 0) \ (\k\{0..<2^(n+1)}. f k) = (\k\{0..< 2^n}. f (2*k+1))" using assms proof(rule nat_induct_at_least) show "\f::(nat \ complex).(\i. i<2^(1+1) \ even i \ f i = 0) \ (\k\{0..<2^(1+1)}. f k) = (\k\{0..< 2^1}. f (2*k+1))" proof(rule allI,rule impI) fix f:: "nat \complex" assume asm: "(\i. i<2^(1+1) \ even i \ f i = 0)" moreover have "(\k\{0..<4}. f k) = f 0 + f 1 + f 2 + f 3" by (simp add: add.commute add.left_commute) moreover have "f 0 = 0" using asm by simp moreover have "f 2 = 0" using asm by simp moreover have "(\k \ {0..< 2^1}. f (2*k+1)) = f 1 + f 3" using add.commute add.left_commute by simp ultimately show "(\k\{0..<2^(1+1)}. f k) = (\k\{0..< 2^1}. f (2*k+1))" by simp qed next fix n:: nat assume "n \ 1" and IH: "\f::(nat \complex).(\i. i<2^(n+1) \ even i \ f i = 0) \ (\k\{0..<2^(n+1)}. f k) = (\k\{0..< 2^n}. f (2*k+1))" show "\f::(nat \complex).(\i. i<2^((Suc n)+1) \ even i \ f i = 0) \ (\k\{0..<2^((Suc n)+1)}. f k) = (\k\{0..< 2^(Suc n)}. f (2*k+1))" proof (rule allI,rule impI) fix f::"nat \complex" assume asm: "(\i. i<2^((Suc n)+1) \ even i \ f i = 0)" have f0: "(\k \{0..<2^(n+1)}. f k) = (\k \ {0..< 2^n}. f (2*k+1))" using asm IH by simp have f1: "(\k\{0..<2^(n+1)}. (\x. f (x+2^(n+1))) k) = (\k\{0..< 2^n}. (\x. f (x+2^(n+1))) (2*k+1))" using asm IH by simp have "(\k\{0..<2^(n+2)}. f k) = (\k\{0..<2^(n+1)}. f k) + (\k\{2^(n+1)..<2^(n+2)}. f k)" by (simp add: sum.atLeastLessThan_concat) also have "... = (\k\{0..< 2^n}. f (2*k+1)) + (\k\{2^(n+1)..<2^(n+2)}. f k)" using f0 by simp also have "... = (\k\{0..< 2^n}. f (2*k+1)) + (\k\{0..<2^(n+1)}. f (k+(2^(n+1))))" using sum.shift_bounds_nat_ivl[of "f" "0" "2^(n+1)" "2^(n+1)"] by simp also have "... = (\k\{0..< 2^n}. f (2*k+1)) + (\k\{0..< 2^n}. (\x. f (x+2^(n+1))) (2*k+1))" using f1 by simp also have "... = (\k\{0..< 2^n}. f (2*k+1)) + (\k\{2^n..< 2^(n+1)}. f (2 *k+1))" using sum.shift_bounds_nat_ivl[of "\x. (f::nat\complex) (2*(x-2^n)+1+2^(n+1))" "0" "2^n" "2^n"] by (simp add: mult_2) also have "... = (\k\{0..< 2^(n+1)}. f (2*k+1))" by (metis Suc_eq_plus1 lessI less_imp_le_nat one_le_numeral power_increasing sum.atLeastLessThan_concat zero_le) finally show "(\k\{0..<2^((Suc n)+1)}. f k) = (\k\{0..< 2^(Suc n)}. f (2*k+1))" by (metis Suc_eq_plus1 add_2_eq_Suc') qed qed lemma (in jozsa) iter_tensor_of_H_times_\\<^sub>2_is_\\<^sub>3: shows "((H^\<^sub>\ n) \ Id 1) * \\<^sub>2 = \\<^sub>3" proof fix i j assume a0:"i < dim_row \\<^sub>3" and a1:"j < dim_col \\<^sub>3" then have f0: "i < (2^(n+1)) \ j = 0" by simp have f1: "((HId^\<^sub>\ n)* \\<^sub>2) $$ (i,j) = (\k<(2^(n+1)). ((HId^\<^sub>\ n) $$ (i,k)) * (\\<^sub>2 $$ (k,j)))" using a1 f0 by (simp add: atLeast0LessThan) show "(((H^\<^sub>\ n) \ Id 1) * \\<^sub>2) $$ (i,j) = \\<^sub>3 $$ (i,j)" proof(rule disjE) show "even i \ odd i" by simp next assume a2: "even i" have "(\(i mod 2 = k mod 2) \ k\ n)) \ ((HId^\<^sub>\ n) $$ (i,k)) * (\\<^sub>2 $$ (k,j)) = 0" for k using f0 by simp then have "k<(2^(n+1)) \ odd k \ ((HId^\<^sub>\ n) $$ (i,k)) * (\\<^sub>2 $$ (k,j)) = 0" for k using a2 mod_2_is_both_even_or_odd f0 by (metis (no_types, lifting) dim_col_mat(1)) then have "(\k\{(0::nat)..<(2^(n+1))}. ((HId^\<^sub>\ n) $$ (i,k)) * (\\<^sub>2 $$ (k,j))) = (\k\{(0::nat)..< (2^n)}. ((HId^\<^sub>\ n) $$ (i,2*k)) * (\\<^sub>2 $$ (2*k,j)))" using sum_every_odd_summand_is_zero dim by simp moreover have "(\k<2^n. ((HId^\<^sub>\ n) $$ (i,2*k)) * (\\<^sub>2 $$ (2*k,j))) = (\k<2^n.(-1)^((i div 2) \\<^bsub>n\<^esub> k)/(sqrt(2)^n) *((-1)^f(k))/(sqrt(2)^(n+1)))" proof- have "(even k \ k\<^sub>2) \ (\\<^sub>2 $$ (k,j)) = ((-1)^f(k div 2))/(sqrt(2)^(n+1))" for k using a0 a1 by simp then have "(\k<2^n. ((HId^\<^sub>\ n) $$ (i,2*k)) * (\\<^sub>2 $$ (2*k,j))) = (\k<2^n. ((HId^\<^sub>\ n) $$ (i,2*k)) *((-1)^f((2*k) div 2))/(sqrt(2)^(n+1)))" by simp moreover have "(even k \ k\ n)) \ ((HId^\<^sub>\ n) $$ (i,k)) = (-1)^ ((i div 2) \\<^bsub>n\<^esub> (k div 2))/(sqrt(2)^n)" for k using a2 a0 a1 by simp ultimately have "(\k<2^n. ((HId^\<^sub>\ n) $$ (i,2*k)) * (\\<^sub>2 $$ (2*k,j))) = (\k<2^n. (-1)^((i div 2) \\<^bsub>n\<^esub> ((2*k) div 2))/(sqrt(2)^n) * ((-1)^f((2*k) div 2))/(sqrt(2)^(n+1)))" by simp then show "(\k<2^n. ((HId^\<^sub>\ n) $$ (i,2*k)) * (\\<^sub>2 $$ (2*k,j))) = (\k<2^n. (-1)^((i div 2) \\<^bsub>n\<^esub> k)/(sqrt(2)^n) *((-1)^f(k))/(sqrt(2)^(n+1)))" by simp qed ultimately have "((HId^\<^sub>\ n)* \\<^sub>2) $$ (i,j) = (\k<2^n. (-1)^((i div 2) \\<^bsub>n\<^esub> k)/(sqrt(2)^n) * ((-1)^f(k))/(sqrt(2)^(n+1)))" using f1 by (metis atLeast0LessThan) also have "... = (\k<2^n. (-1)^(f(k)+((i div 2) \\<^bsub>n\<^esub> k))/((sqrt(2)^n)*(sqrt(2)^(n+1))))" by (simp add: power_add mult.commute) finally have "((HId^\<^sub>\ n)* \\<^sub>2) $$ (i,j) = (\k<2^n. (-1)^(f(k)+((i div 2) \\<^bsub>n\<^esub> k))/((sqrt(2)^n)*(sqrt(2)^(n+1))))" by simp moreover have "\\<^sub>3 $$ (i,j) = (\k<2^n. (-1)^(f(k) + ((i div 2) \\<^bsub>n\<^esub> k))/(sqrt(2)^n * sqrt(2)^(n+1)))" using a0 a1 a2 by simp ultimately show "(((H^\<^sub>\ n) \ Id 1)* \\<^sub>2) $$ (i,j) = \\<^sub>3 $$ (i,j)" using iter_tensor_of_H_tensor_Id_is_HId dim by simp next assume a2: "odd i" have "(\(i mod 2 = k mod 2) \ k\ n)) \ ((HId^\<^sub>\ n) $$ (i,k)) * (\\<^sub>2 $$ (k,j)) = 0" for k using f0 by simp then have "k<(2^(n+1)) \ even k \ ((HId^\<^sub>\ n) $$ (i,k)) * (\\<^sub>2 $$ (k,j)) = 0" for k using a2 mod_2_is_both_even_or_odd f0 by (metis (no_types, lifting) dim_col_mat(1)) then have "(\k\{0..<2^(n+1)}. ((HId^\<^sub>\ n) $$ (i,k)) * (\\<^sub>2 $$ (k,j))) = (\k\{0..<2^n}. ((HId^\<^sub>\ n) $$ (i,2*k+1)) * (\\<^sub>2 $$ (2*k+1,j)))" using sum_every_even_summand_is_zero dim by simp moreover have "(\k<2^n. ((HId^\<^sub>\ n) $$ (i,2*k+1)) * (\\<^sub>2 $$ (2*k+1,j))) = (\ k<2^n. (-1)^((i div 2) \\<^bsub>n\<^esub> k)/(sqrt(2)^n) * ((-1)^(f(k)+1))/(sqrt(2)^(n+1)))" proof- have "(odd k \ k\<^sub>2) \ (\\<^sub>2 $$ (k,j)) = ((-1)^(f(k div 2)+1))/(sqrt(2)^(n+1))" for k using a0 a1 a2 by simp then have f2:"(\k<2^n. ((HId^\<^sub>\ n) $$ (i,2*k+1)) * (\\<^sub>2 $$ (2*k+1,j))) = (\k<2^n. ((HId^\<^sub>\ n) $$ (i,2*k+1)) * ((-1)^(f((2*k+1) div 2)+1))/(sqrt(2)^(n+1)))" by simp have "i < dim_row (HId^\<^sub>\ n)" using f0 a2 mod_2_is_both_even_or_odd by simp then have "((i mod 2 = k mod 2) \ k\ n)) \ ((HId^\<^sub>\ n) $$ (i,k)) = (-1)^((i div 2) \\<^bsub>n\<^esub> (k div 2))/(sqrt(2)^n) " for k using a2 a0 a1 f0 dim HId_values by simp moreover have "odd k \ (i mod 2 = k mod 2)" for k using a2 mod_2_is_both_even_or_odd by auto ultimately have "(odd k \ k\ n)) \ ((HId^\<^sub>\ n) $$ (i,k)) = (-1)^((i div 2) \\<^bsub>n\<^esub> (k div 2))/(sqrt(2)^n)" for k by simp then have "k<2^n \ ((HId^\<^sub>\ n) $$ (i,2*k+1)) = (-1)^((i div 2) \\<^bsub>n\<^esub> ((2*k+1) div 2))/(sqrt(2)^n) " for k by simp then have "(\k<2^n. ((HId^\<^sub>\ n) $$ (i,2*k+1)) * (\\<^sub>2 $$ (2*k+1,j))) = (\k<2^n. (-1)^((i div 2) \\<^bsub>n\<^esub> ((2*k+1) div 2))/(sqrt(2)^n) * ((-1)^(f((2*k+1) div 2)+1))/(sqrt(2)^(n+1)))" using f2 by simp then show "(\k<2^n. ((HId^\<^sub>\ n) $$ (i,2*k+1)) * (\\<^sub>2 $$ (2*k+1,j))) = (\k<2^n. (-1)^((i div 2) \\<^bsub>n\<^esub> k)/(sqrt(2)^n) *((-1)^(f(k)+1))/(sqrt(2)^(n+1)))" by simp qed ultimately have "((HId^\<^sub>\ n)* \\<^sub>2) $$ (i,j) = (\k<2^n. (-1)^((i div 2) \\<^bsub>n\<^esub> k)/(sqrt(2)^n) * ((-1)^(f(k)+1))/(sqrt(2)^(n+1)))" using f1 by (metis atLeast0LessThan) also have "... = (\k<2^n. (-1)^(f(k)+1+((i div 2) \\<^bsub>n\<^esub> k))/((sqrt(2)^n)*(sqrt(2)^(n+1))))" by (simp add: mult.commute power_add) finally have "((HId^\<^sub>\ n)* \\<^sub>2) $$ (i,j) = (\k< 2^n. (-1)^(f(k)+1+((i div 2) \\<^bsub>n\<^esub> k))/((sqrt(2)^n)*(sqrt(2)^(n+1))))" by simp then show "(((H^\<^sub>\ n) \ Id 1)* \\<^sub>2) $$ (i,j) = \\<^sub>3 $$ (i,j)" using iter_tensor_of_H_tensor_Id_is_HId dim a2 a0 a1 by simp qed next show "dim_row (((H^\<^sub>\ n) \ Id 1) * \\<^sub>2) = dim_row \\<^sub>3" using iter_tensor_of_H_tensor_Id_is_HId dim by simp next show "dim_col (((H^\<^sub>\ n) \ Id 1)* \\<^sub>2) = dim_col \\<^sub>3" using iter_tensor_of_H_tensor_Id_is_HId dim by simp qed lemma (in jozsa) \\<^sub>3_is_state: shows "state (n+1) \\<^sub>3" proof- have "((H^\<^sub>\ n) \ Id 1) * \\<^sub>2 = \\<^sub>3" using iter_tensor_of_H_times_\\<^sub>2_is_\\<^sub>3 by simp moreover have "gate (n+1) ((H^\<^sub>\ n) \ Id 1)" using iter_tensor_of_H_tensor_Id_is_HId HId_is_gate dim by simp moreover have "state (n+1) \\<^sub>2" using \\<^sub>2_is_state by simp ultimately show "state (n+1) \\<^sub>3" using gate_on_state_is_state dim by (metis (no_types, lifting)) qed text \ Finally, all steps are put together. The result depends on the function f. If f is constant the first n qubits are 0, if f is balanced there is at least one qubit in state 1 among the first n qubits. The algorithm only uses one evaluation of f(x) and will always succeed. \ definition (in jozsa) jozsa_algo:: "complex Matrix.mat" where "jozsa_algo \ ((H \\<^bsup>n\<^esup>) \ Id 1) * (U\<^sub>f * (((H \\<^bsup>n\<^esup>) * ( |zero\ \\<^bsup>n\<^esup>)) \ (H * |one\)))" lemma (in jozsa) jozsa_algo_result [simp]: shows "jozsa_algo = \\<^sub>3" using jozsa_algo_def H_on_ket_one_is_\\<^sub>1\<^sub>1 iter_tensor_of_H_on_zero_tensor \\<^sub>1\<^sub>0_tensor_\\<^sub>1\<^sub>1_is_\\<^sub>1 jozsa_transform_times_\\<^sub>1_is_\\<^sub>2 iter_tensor_of_H_times_\\<^sub>2_is_\\<^sub>3 dim iter_tensor_of_H_rep_is_correct by simp lemma (in jozsa) jozsa_algo_result_is_state: shows "state (n+1) jozsa_algo" using \\<^sub>3_is_state by simp lemma (in jozsa) prob0_fst_qubits_of_jozsa_algo: shows "(prob0_fst_qubits n jozsa_algo) = (\j\{0,1}. (cmod(jozsa_algo $$ (j,0)))\<^sup>2)" using prob0_fst_qubits_eq by simp text \General lemmata required to compute probabilities.\ lemma aux_comp_with_sqrt2: shows "(sqrt 2)^n * (sqrt 2)^n = 2^n" - by (smt power_mult_distrib real_sqrt_mult_self) + by (smt (verit) power_mult_distrib real_sqrt_mult_self) lemma aux_comp_with_sqrt2_bis [simp]: shows "2^n/(sqrt(2)^n * sqrt(2)^(n+1)) = 1/sqrt 2" using aux_comp_with_sqrt2 by (simp add: mult.left_commute) lemma aux_ineq_with_card: fixes g:: "nat \ nat" and A:: "nat set" assumes "finite A" shows "(\k\A. (-1)^(g k)) \ card A" and "(\k\A. (-1)^(g k)) \ -card A" - apply (smt assms neg_one_even_power neg_one_odd_power card_eq_sum of_nat_1 of_nat_sum sum_mono) - apply (smt assms neg_one_even_power neg_one_odd_power card_eq_sum of_nat_1 of_nat_sum sum_mono sum_negf). + apply (smt (verit) assms neg_one_even_power neg_one_odd_power card_eq_sum of_nat_1 of_nat_sum sum_mono) + apply (smt (verit) assms neg_one_even_power neg_one_odd_power card_eq_sum of_nat_1 of_nat_sum sum_mono sum_negf). lemma aux_comp_with_cmod: fixes g:: "nat \ nat" assumes "(\x<2^n. g x = 0) \ (\x<2^n. g x = 1)" shows "(cmod (\k<2^n. (-1)^(g k)))\<^sup>2 = 2^(2*n)" proof(rule disjE) show "(\x<2^n. g x = 0) \ (\x<2^n. g x = 1)" using assms by simp next assume "\x<2^n. g x = 0" then have "(cmod (\k<2^n. (-1)^(g k)))\<^sup>2 = (2^n)\<^sup>2" by (simp add: norm_power) then show "?thesis" by (simp add: power_even_eq) next assume "\x<2^n. g x = 1" then have "(cmod (\k<2^n. (-1)^(g k)))\<^sup>2 = (2^n)\<^sup>2" by (simp add: norm_power) then show "?thesis" by (simp add: power_even_eq) qed lemma cmod_less: fixes a n:: int assumes "a < n" and "a > -n" shows "cmod a < n" using assms by simp lemma square_less: fixes a n:: real assumes "a < n" and "a > -n" shows "a\<^sup>2 < n\<^sup>2" - using assms by (smt power2_eq_iff power2_minus power_less_imp_less_base) + using assms by (smt (verit) power2_eq_iff power2_minus power_less_imp_less_base) lemma cmod_square_real [simp]: fixes n:: real shows "(cmod n)\<^sup>2 = n\<^sup>2" by simp lemma aux_comp_sum_divide_cmod: fixes n:: nat and g:: "nat \ int" and a:: real shows "(cmod(complex_of_real(\k2 = (cmod (\k2" by (metis cmod_square_real of_int_sum of_real_of_int_eq power_divide sum_divide_distrib) text \ The function is constant if and only if the first n qubits are 0. So, if the function is constant, then the probability of measuring 0 for the first n qubits is 1. \ lemma (in jozsa) prob0_jozsa_algo_of_const_0: assumes "const 0" shows "prob0_fst_qubits n jozsa_algo = 1" proof- have "prob0_fst_qubits n jozsa_algo = (\j\{0,1}. (cmod(jozsa_algo $$ (j,0)))\<^sup>2)" using prob0_fst_qubits_of_jozsa_algo by simp moreover have "(cmod(jozsa_algo $$ (0,0)))\<^sup>2 = 1/2" proof- have "k<2^n \ ((0 div 2) \\<^bsub>n\<^esub> k) = 0" for k::nat using bitwise_inner_prod_with_zero by simp then have "(cmod(jozsa_algo $$ (0,0)))\<^sup>2 = (cmod(\k::nat<2^n. 1/(sqrt(2)^n * sqrt(2)^(n+1))))\<^sup>2" using jozsa_algo_result const_def assms by simp also have "... = (cmod((2::nat)^n/(sqrt(2)^n * sqrt(2)^(n+1))))\<^sup>2" by simp also have "... = (cmod(1/(sqrt(2))))\<^sup>2" using aux_comp_with_sqrt2_bis by simp also have "... = 1/2" by (simp add: norm_divide power2_eq_square) finally show "?thesis" by simp qed moreover have "(cmod(jozsa_algo $$ (1,0)))\<^sup>2 = 1/2" proof- have "k<2^n \ ((1 div 2) \\<^bsub>n\<^esub> k) = 0" for k:: nat using bitwise_inner_prod_with_zero by simp then have "k<2^n \ f k + 1 + ((1 div 2) \\<^bsub>n\<^esub> k) = 1" for k::nat using const_def assms by simp moreover have "(cmod(jozsa_algo $$ (1,0)))\<^sup>2 = (cmod (\k::nat<2^n. (-1)^(f k + 1 + ((1 div 2) \\<^bsub>n\<^esub> k))/(sqrt(2)^n * sqrt(2)^(n+1))))\<^sup>2" using \\<^sub>3_dim by simp ultimately have "(cmod(jozsa_algo $$ (1,0)))\<^sup>2 = (cmod(\k::nat<2^n. -1/(sqrt(2)^n * sqrt(2)^(n+1))))\<^sup>2" - by (smt lessThan_iff power_one_right sum.cong) + by (smt (verit) lessThan_iff power_one_right sum.cong) also have "... = (cmod(-1/(sqrt(2))))\<^sup>2" using aux_comp_with_sqrt2_bis by simp also have "... = 1/2" by (simp add: norm_divide power2_eq_square) finally show "?thesis" by simp qed ultimately have "prob0_fst_qubits n jozsa_algo = 1/2 + 1/2" by simp then show "?thesis" by simp qed lemma (in jozsa) prob0_jozsa_algo_of_const_1: assumes "const 1" shows "prob0_fst_qubits n jozsa_algo = 1" proof- have "prob0_fst_qubits n jozsa_algo = (\j\{0,1}. (cmod(jozsa_algo $$ (j,0)))\<^sup>2)" using prob0_fst_qubits_of_jozsa_algo by simp moreover have "(cmod(jozsa_algo $$ (0,0)))\<^sup>2 = 1/2" proof- have "k<2^n \ ((0 div 2) \\<^bsub>n\<^esub> k) = 0" for k::nat using bitwise_inner_prod_with_zero by simp then have "(cmod(jozsa_algo $$ (0,0)))\<^sup>2 = (cmod(\k::nat<2^n. 1/(sqrt(2)^n * sqrt(2)^(n+1))))\<^sup>2" using jozsa_algo_result const_def assms by simp also have "... = (cmod((-1)/(sqrt(2))))\<^sup>2 " using aux_comp_with_sqrt2_bis by simp also have "... = 1/2" by (simp add: norm_divide power2_eq_square) finally show "?thesis" by simp qed moreover have "(cmod(jozsa_algo $$ (1,0)))\<^sup>2 = 1/2" proof- have "k<2^n \ ((1 div 2) \\<^bsub>n\<^esub> k) = 0" for k::nat using bitwise_inner_prod_with_zero by simp then have "(\k::nat<2^n. (-1)^(f k +1 + ((1 div 2) \\<^bsub>n\<^esub> k))/(sqrt(2)^n * sqrt(2)^(n+1))) = (\k::nat<2^n. 1/(sqrt(2)^n * sqrt(2)^(n+1)))" using const_def assms by simp moreover have "(cmod(jozsa_algo $$ (1,0)))\<^sup>2 = (cmod (\k::nat<2^n. (-1)^(f k + 1 + ((1 div 2) \\<^bsub>n\<^esub> k))/(sqrt(2)^n * sqrt(2)^(n+1))))\<^sup>2" using \\<^sub>3_dim by simp ultimately have "(cmod(jozsa_algo $$ (1,0)))\<^sup>2 = (cmod(\k::nat<2^n. 1/(sqrt(2)^n * sqrt(2)^(n+1))))\<^sup>2" by simp also have "... = (cmod(1/(sqrt(2))))\<^sup>2 " using aux_comp_with_sqrt2_bis by simp also have "... = 1/2" by (simp add: norm_divide power2_eq_square) finally show "?thesis" by simp qed ultimately have "prob0_fst_qubits n jozsa_algo = 1/2 + 1/2" by simp then show "?thesis" by simp qed text \If the probability of measuring 0 for the first n qubits is 1, then the function is constant.\ lemma (in jozsa) max_value_of_not_const_less: assumes "\ const 0" and "\ const 1" shows "(cmod (\k::nat<2^n. (-(1::nat))^(f k)))\<^sup>2 < (2::nat)^(2*n)" proof- have "cmod (\k::nat<2^n. (-(1::nat))^(f k)) < 2^n" proof- have "(\k::nat<2^n. (-(1::nat))^(f k)) < 2^n" proof- obtain x where f0:"x < 2^n" and f1:"f x = 1" using assms(1) const_def f_values by auto then have "(\k::nat<2^n. (-(1::nat))^(f k)) < (\k\{i| i:: nat. i < 2^n}-{x}. (-(1::nat))^(f k))" proof- have "(-(1::nat))^ f x = -1" using f1 by simp moreover have "x\{i| i::nat. i<2^n}" using f0 by simp moreover have "finite {i| i::nat. i<2^n}" by simp moreover have "(\k\{i| i::nat. i<2^n}. (-(1::nat))^(f k)) < (\k\{i| i:: nat. i<2^n}-{x}. (-(1::nat))^(f k))" using calculation(1,2,3) sum_diff1 by (simp add: sum_diff1) ultimately show ?thesis by (metis Collect_cong Collect_mem_eq lessThan_iff) qed moreover have "\ \ int (2^n - 1)" using aux_ineq_with_card(1)[of "{i| i:: nat. i<2^n}-{x}"] f0 by simp ultimately show ?thesis by (meson diff_le_self less_le_trans of_nat_le_numeral_power_cancel_iff) qed moreover have "(\k::nat<2^n. (-(1::nat))^(f k)) > - (2^n)" proof- obtain x where f0:"x < 2^n" and f1:"f x = 0" using assms(2) const_def f_values by auto then have "(\k::nat<2^n. (-(1::nat))^(f k)) > (\k\{i| i:: nat. i < 2^n}-{x}. (-(1::nat))^(f k))" proof- have "(-(1::nat))^ f x = 1" using f1 by simp moreover have "x\{i| i::nat. i<2^n}" using f0 by simp moreover have "finite {i| i::nat. i<2^n}" by simp moreover have "(\k\{i| i::nat. i<2^n}. (-(1::nat))^(f k)) > (\k\{i| i:: nat. i<2^n}-{x}. (-(1::nat))^(f k))" using calculation(1,2,3) sum_diff1 by (simp add: sum_diff1) ultimately show ?thesis by (metis Collect_cong Collect_mem_eq lessThan_iff) qed moreover have "- int (2^n - 1) \ (\k\{i| i:: nat. i < 2^n}-{x}. (-(1::nat))^(f k))" using aux_ineq_with_card(2)[of "{i| i:: nat. i<2^n}-{x}"] f0 by simp ultimately show ?thesis - by (smt diff_le_self of_nat_1 of_nat_add of_nat_power_le_of_nat_cancel_iff one_add_one) + by (smt (verit) diff_le_self of_nat_1 of_nat_add of_nat_power_le_of_nat_cancel_iff one_add_one) qed ultimately show ?thesis using cmod_less of_int_of_nat_eq of_nat_numeral of_nat_power by (metis (no_types, lifting)) qed then have "(cmod (\k::nat<2^n. (-(1::nat))^(f k)))\<^sup>2 < (2^n)\<^sup>2" - using square_less norm_ge_zero by smt + using square_less norm_ge_zero by (smt (verit)) thus ?thesis by (simp add: power_even_eq) qed lemma (in jozsa) max_value_of_not_const_less_bis: assumes "\ const 0" and "\ const 1" shows "(cmod (\k::nat<2^n. (-(1::nat))^(f k + 1)))\<^sup>2 < (2::nat)^(2*n)" proof- have "cmod (\k::nat<2^n. (-(1::nat))^(f k + 1)) < 2^n" proof- have "(\k::nat<2^n. (-(1::nat))^(f k + 1)) < 2^n" proof- obtain x where f0:"x < 2^n" and f1:"f x = 0" using assms(2) const_def f_values by auto then have "(\k::nat<2^n. (-(1::nat))^(f k + 1)) < (\k\{i| i:: nat. i < 2^n}-{x}. (-(1::nat))^(f k + 1))" proof- have "(-(1::nat))^ (f x + 1) = -1" using f1 by simp moreover have "x\{i| i::nat. i<2^n}" using f0 by simp moreover have "finite {i| i::nat. i<2^n}" by simp moreover have "(\k\{i| i::nat. i<2^n}. (-(1::nat))^(f k + 1)) < (\k\{i| i:: nat. i<2^n}-{x}. (-(1::nat))^(f k + 1))" using calculation(1,2,3) sum_diff1 by (simp add: sum_diff1) ultimately show ?thesis by (metis Collect_cong Collect_mem_eq lessThan_iff) qed moreover have "\ \ int (2^n - 1)" using aux_ineq_with_card(1)[of "{i| i:: nat. i<2^n}-{x}" "\k. f k + 1"] f0 by simp ultimately show ?thesis by (meson diff_le_self less_le_trans of_nat_le_numeral_power_cancel_iff) qed moreover have "(\k::nat<2^n. (-(1::nat))^(f k + 1)) > - (2^n)" proof- obtain x where f0:"x < 2^n" and f1:"f x = 1" using assms(1) const_def f_values by auto then have "(\k::nat<2^n. (-(1::nat))^(f k + 1)) > (\k\{i| i:: nat. i < 2^n}-{x}. (-(1::nat))^(f k + 1))" proof- have "(-(1::nat))^ (f x + 1) = 1" using f1 by simp moreover have "x\{i| i::nat. i<2^n}" using f0 by simp moreover have "finite {i| i::nat. i<2^n}" by simp moreover have "(\k\{i| i::nat. i<2^n}. (-(1::nat))^(f k + 1)) > (\k\{i| i:: nat. i<2^n}-{x}. (-(1::nat))^(f k + 1))" using calculation(1,2,3) sum_diff1 by (simp add: sum_diff1) ultimately show ?thesis by (metis Collect_cong Collect_mem_eq lessThan_iff) qed moreover have "- int (2^n - 1) \ (\k\{i| i:: nat. i < 2^n}-{x}. (-(1::nat))^(f k + 1))" using aux_ineq_with_card(2)[of "{i| i:: nat. i<2^n}-{x}" "\k. f k + 1"] f0 by simp ultimately show ?thesis - by (smt diff_le_self of_nat_1 of_nat_add of_nat_power_le_of_nat_cancel_iff one_add_one) + by (smt (verit) diff_le_self of_nat_1 of_nat_add of_nat_power_le_of_nat_cancel_iff one_add_one) qed ultimately show ?thesis using cmod_less of_int_of_nat_eq of_nat_numeral of_nat_power by (metis (no_types, lifting)) qed then have "(cmod (\k::nat<2^n. (-(1::nat))^(f k + 1)))\<^sup>2 < (2^n)\<^sup>2" - using square_less norm_ge_zero by smt + using square_less norm_ge_zero by (smt (verit)) thus ?thesis by (simp add: power_even_eq) qed lemma (in jozsa) f_const_has_max_value: assumes "const 0 \ const 1" shows "(cmod (\k<(2::nat)^n. (-1)^(f k)))\<^sup>2 = (2::nat)^(2*n)" and "(cmod (\k<(2::nat)^n. (-1)^(f k + 1)))\<^sup>2 = (2::nat)^(2*n)" using aux_comp_with_cmod[of n "\k. f k"] aux_comp_with_cmod[of n "\k. f k + 1"] const_def assms by auto lemma (in jozsa) prob0_fst_qubits_leq: shows "(cmod (\k<(2::nat)^n. (-1)^(f k)))\<^sup>2 \ (2::nat)^(2*n)" and "(cmod (\k<(2::nat)^n. (-1)^(f k + 1)))\<^sup>2 \ (2::nat)^(2*n)" proof- show "(cmod (\k<(2::nat)^n. (-1)^(f k)))\<^sup>2 \ (2::nat)^(2*n)" proof(rule disjE) show "(const 0 \ const 1) \ (\ const 0 \ \ const 1)" by auto next assume "const 0 \ const 1" then show "(cmod (\k<(2::nat)^n. (-1)^(f k)))\<^sup>2 \ (2::nat)^(2*n)" using f_const_has_max_value by simp next assume "\ const 0 \ \ const 1" then show "(cmod (\k<(2::nat)^n. (-1)^(f k)))\<^sup>2 \ (2::nat)^(2*n)" using max_value_of_not_const_less by simp qed next show "(cmod (\k<(2::nat)^n. (-1)^(f k + 1)))\<^sup>2 \ (2::nat)^(2*n)" proof(rule disjE) show "(const 0 \ const 1) \ (\ const 0 \ \ const 1)" by auto next assume "const 0 \ const 1" then show "(cmod (\k<(2::nat)^n. (-1)^(f k + 1)))\<^sup>2 \ (2::nat)^(2*n)" using f_const_has_max_value by simp next assume "\ const 0 \ \ const 1" then show "(cmod (\k<(2::nat)^n. (-1)^(f k + 1)))\<^sup>2 \ (2::nat)^(2*n)" using max_value_of_not_const_less_bis by simp qed qed lemma (in jozsa) prob0_jozsa_algo_1_is_const: assumes "prob0_fst_qubits n jozsa_algo = 1" shows "const 0 \ const 1" proof- have f0: "(\j\{0,1}. (cmod(jozsa_algo $$ (j,0)))\<^sup>2) = 1" using prob0_fst_qubits_of_jozsa_algo assms by simp have "k < 2^n\((0 div 2) \\<^bsub>n\<^esub> k) = 0" for k::nat using bitwise_inner_prod_with_zero by simp then have f1: "(cmod(jozsa_algo $$ (0,0)))\<^sup>2 = (cmod(\k<(2::nat)^n. (-1)^(f k)/(sqrt(2)^n * sqrt(2)^(n+1))))\<^sup>2" by simp have "k < 2^n\((1 div 2) \\<^bsub>n\<^esub> k) = 0" for k::nat using bitwise_inner_prod_with_zero by simp moreover have "(cmod(jozsa_algo $$ (1,0)))\<^sup>2 = (cmod (\k<(2::nat)^n. (-1)^(f k+ 1 + ((1 div 2) \\<^bsub>n\<^esub> k))/(sqrt(2)^n * sqrt(2)^(n+1))))\<^sup>2" using \\<^sub>3_dim by simp ultimately have f2: "(cmod(jozsa_algo $$ (1,0)))\<^sup>2 = (cmod (\k<(2::nat)^n. (-1)^(f k + 1)/(sqrt(2)^n * sqrt(2)^(n+1))))\<^sup>2" by simp have f3: "1 = (cmod(\k::nat<(2::nat)^n.(-1)^(f k)/(sqrt(2)^n * sqrt(2)^(n+1))))\<^sup>2 + (cmod (\k::nat<(2::nat)^n. (-1)^(f k + 1)/(sqrt(2)^n * sqrt(2)^(n+1))))\<^sup>2" using f0 f1 f2 by simp also have "... = ((cmod (\k::nat<(2::nat)^n. (-1)^(f k)) ) /(sqrt(2)^n * sqrt(2)^(n+1)))\<^sup>2 + ((cmod(\k::nat<(2::nat)^n. (-1)^(f k + 1))) /(sqrt(2)^n * sqrt(2)^(n+1)))\<^sup>2" using aux_comp_sum_divide_cmod[of "\k. (-1)^(f k)" "(sqrt(2)^n * sqrt(2)^(n+1))" "(2::nat)^n"] aux_comp_sum_divide_cmod[of "\k. (-1)^(f k + 1)" "(sqrt(2)^n * sqrt(2)^(n+1))" "(2::nat)^n"] by simp also have "... = ((cmod (\k::nat<(2::nat)^n. (-1)^(f k))))\<^sup>2 /((sqrt(2)^n * sqrt(2)^(n+1)))\<^sup>2 + ((cmod(\k::nat<(2::nat)^n. (-1)^(f k +1))))\<^sup>2 /((sqrt(2)^n * sqrt(2)^(n+1)))\<^sup>2" by (simp add: power_divide) also have "... = ((cmod (\k::nat<(2::nat)^n. (-1)^(f k)) ) )\<^sup>2/(2^(2*n+1)) + ((cmod(\k::nat<(2::nat)^n. (-1)^(f k + 1))))\<^sup>2 /(2^(2*n+1))" - by (smt left_add_twice power2_eq_square power_add power_mult_distrib real_sqrt_pow2) + by (smt (verit) left_add_twice power2_eq_square power_add power_mult_distrib real_sqrt_pow2) also have "... = (((cmod (\k::nat<(2::nat)^n. (-1)^(f k))))\<^sup>2 + ((cmod(\k::nat<(2::nat)^n. (-1)^(f k + 1))))\<^sup>2)/(2^(2*n+1)) " by (simp add: add_divide_distrib) finally have "((2::nat)^(2*n+1)) = (((cmod (\k::nat<(2::nat)^n. (-1)^(f k))))\<^sup>2 + ((cmod(\k::nat<(2::nat)^n. (-1)^(f k + 1))))\<^sup>2)" by simp moreover have "((2::nat)^(2*n+1)) = 2^(2*n) + 2^(2*n)" by auto moreover have "(cmod (\k<(2::nat)^n. (-1)^(f k)))\<^sup>2 \ 2^(2*n)" using prob0_fst_qubits_leq by simp moreover have "(cmod (\k<(2::nat)^n. (-1)^(f k + 1)))\<^sup>2 \ 2^(2*n)" using prob0_fst_qubits_leq by simp ultimately have "2^(2*n) = ((cmod (\k::nat<(2::nat)^n. (-1)^(f k))))\<^sup>2" by simp then show ?thesis using max_value_of_not_const_less by auto qed text \ The function is balanced if and only if at least one qubit among the first n qubits is not zero. So, if the function is balanced then the probability of measuring 0 for the first n qubits is 0. \ lemma sum_union_disjoint_finite_set: fixes C::"nat set" and g::"nat \ int" assumes "finite C" shows "\A B. A \ B = {} \ A \ B = C \ (\k\C. g k) = (\k\A. g k) + (\k\B. g k)" using assms sum.union_disjoint by auto lemma (in jozsa) balanced_pos_and_neg_terms_cancel_out1: assumes "is_balanced" shows "(\k<(2::nat)^n. (-(1::nat))^(f k)) = 0" proof- have "\A B. A \ {i::nat. i < (2::nat)^n} \ B \ {i::nat. i < (2::nat)^n} \ card A = ((2::nat)^(n-1)) \ card B = ((2::nat)^(n-1)) \ (\(x::nat) \ A. f x = (0::nat)) \ (\(x::nat) \ B. f x = 1) \ (\k<(2::nat)^n. (-(1::nat))^(f k)) = 0" proof fix A B::"nat set" assume asm: "A \ {i::nat. i < (2::nat)^n} \ B \ {i::nat. i < (2::nat)^n} \ card A = ((2::nat)^(n-1)) \ card B = ((2::nat)^(n-1)) \ (\(x::nat) \ A. f x = (0::nat)) \ (\(x::nat) \ B. f x = 1)" then have " A \ B = {}" and "{0..<(2::nat)^n} = A \ B" using is_balanced_union is_balanced_inter by auto then have "(\k\{0..<(2::nat)^n}. (-(1::nat))^(f k)) = (\k\A. (-(1::nat))^(f k)) + (\k\B. (-(1::nat))^(f k))" by (metis finite_atLeastLessThan sum_union_disjoint_finite_set) moreover have "(\k\A. (-1)^(f k)) = ((2::nat)^(n-1))" using asm by simp moreover have "(\k\B. (-1)^(f k)) = -((2::nat)^(n-1))" using asm by simp ultimately have "(\k\ {0..<(2::nat)^n}. (-(1::nat))^(f k)) = 0" by simp then show "(\k<(2::nat)^n. (-(1::nat))^(f k)) = 0" by (simp add: lessThan_atLeast0) qed then show ?thesis using assms is_balanced_def by auto qed lemma (in jozsa) balanced_pos_and_neg_terms_cancel_out2: assumes "is_balanced" shows "(\k<(2::nat)^n. (-(1::nat))^(f k + 1)) = 0" proof- have "\A B. A \ {i::nat. i < (2::nat)^n} \ B \ {i::nat. i < (2::nat)^n} \ card A = ((2::nat)^(n-1)) \ card B = ((2::nat)^(n-1)) \ (\(x::nat)\A. f x = (0::nat)) \ (\(x::nat)\B. f x = 1) \ (\k<(2::nat)^n. (-(1::nat))^(f k + 1)) = 0" proof fix A B::"nat set" assume asm: "A \ {i::nat. i < (2::nat)^n} \ B \ {i::nat. i < (2::nat)^n} \ card A = ((2::nat)^(n-1)) \ card B = ((2::nat)^(n-1)) \ (\(x::nat) \ A. f x = (0::nat)) \ (\(x::nat) \ B. f x = 1)" have "A \ B = {}" and "{0..<(2::nat)^n} = A \ B" using is_balanced_union is_balanced_inter asm by auto then have "(\k\{0..<(2::nat)^n}. (-(1::nat))^(f k + 1)) = (\k\A. (-(1::nat))^(f k + 1)) + (\k\B. (-(1::nat))^(f k + 1))" by (metis finite_atLeastLessThan sum_union_disjoint_finite_set) moreover have "(\k\A. (-1)^(f k + 1)) = -((2::nat)^(n-1))" using asm by simp moreover have "(\k\B. (-1)^(f k + 1)) = ((2::nat)^(n-1))" using asm by simp ultimately have "(\k\{0..<(2::nat)^n}. (-(1::nat))^(f k + 1)) = 0 " by simp then show "(\k<(2::nat)^n. (-(1::nat))^(f k + 1)) = 0" by (simp add: lessThan_atLeast0) qed then show "(\k<(2::nat)^n. (-(1::nat))^(f k + 1)) = 0" using assms is_balanced_def by auto qed lemma (in jozsa) prob0_jozsa_algo_of_balanced: assumes "is_balanced" shows "prob0_fst_qubits n jozsa_algo = 0" proof- have "prob0_fst_qubits n jozsa_algo = (\j\{0,1}. (cmod(jozsa_algo $$ (j,0)))\<^sup>2)" using prob0_fst_qubits_of_jozsa_algo by simp moreover have "(cmod(jozsa_algo $$ (0,0)))\<^sup>2 = 0" proof- have "k < 2^n\((1 div 2) \\<^bsub>n\<^esub> k) = 0" for k::nat using bitwise_inner_prod_with_zero by simp then have "(cmod(jozsa_algo $$ (0,0)))\<^sup>2 = (cmod(\ k < (2::nat)^n. (-1)^(f k)/(sqrt(2)^n * sqrt(2)^(n+1))))\<^sup>2" using \\<^sub>3_values by simp also have "... = (cmod(\k<(2::nat)^n. (-(1::nat))^(f k))/(sqrt(2)^n * sqrt(2)^(n+1)))\<^sup>2" using aux_comp_sum_divide_cmod[of "\k.(-(1::nat))^(f k)" "(sqrt(2)^n * sqrt(2)^(n+1))" "2^n"] by simp also have "... = (cmod ((0::int)/(sqrt(2)^n * sqrt(2)^(n+1))))\<^sup>2" using balanced_pos_and_neg_terms_cancel_out1 assms by (simp add: bob_fun_axioms) also have "... = 0" by simp finally show ?thesis by simp qed moreover have "(cmod(jozsa_algo $$ (1,0)))\<^sup>2 = 0" proof- have "k < 2^n \ (((1::nat) div 2) \\<^bsub>n\<^esub> k) = 0" for k::nat using bitwise_inner_prod_with_zero by auto moreover have "(cmod(jozsa_algo $$ (1,0)))\<^sup>2 = (cmod (\k<(2::nat)^n. (-(1::nat))^(f k + (1::nat) + ((1 div 2) \\<^bsub>n\<^esub> k))/(sqrt(2)^n * sqrt(2)^(n+1))))\<^sup>2" using \\<^sub>3_dim by simp ultimately have "(cmod(jozsa_algo $$ (1,0)))\<^sup>2 = (cmod(\k<(2::nat)^n. (-(1::nat))^(f k + (1::nat))/(sqrt(2)^n * sqrt(2)^(n+1))))\<^sup>2" by simp also have "... = (cmod(\k<(2::nat)^n. (-(1::nat))^(f k + 1))/(sqrt(2)^n * sqrt(2)^(n+1)))\<^sup>2" using aux_comp_sum_divide_cmod[of "\k.(-(1::nat))^(f k + 1)" "(sqrt(2)^n * sqrt(2)^(n+1))" "2^n"] by simp also have "... = (cmod ((0::int)/(sqrt(2)^n * sqrt(2)^(n+1))))\<^sup>2" using balanced_pos_and_neg_terms_cancel_out2 assms by (simp add: bob_fun_axioms) also have "... = 0" by simp finally show ?thesis by simp qed ultimately have "prob0_fst_qubits n jozsa_algo = 0 + 0" by simp then show ?thesis by simp qed text \If the probability that the first n qubits are 0 is 0, then the function is balanced.\ lemma (in jozsa) balanced_prob0_jozsa_algo: assumes "prob0_fst_qubits n jozsa_algo = 0" shows "is_balanced" proof- have "is_const \ is_balanced" using const_or_balanced by simp moreover have "is_const \ \ prob0_fst_qubits n jozsa_algo = 0" using is_const_def prob0_jozsa_algo_of_const_0 prob0_jozsa_algo_of_const_1 by simp ultimately show ?thesis using assms by simp qed text \We prove the correctness of the algorithm.\ definition (in jozsa) jozsa_algo_eval:: "real" where "jozsa_algo_eval \ prob0_fst_qubits n jozsa_algo" theorem (in jozsa) jozsa_algo_is_correct: shows "jozsa_algo_eval = 1 \ is_const" and "jozsa_algo_eval = 0 \ is_balanced" using prob0_jozsa_algo_of_const_1 prob0_jozsa_algo_of_const_0 jozsa_algo_eval_def prob0_jozsa_algo_1_is_const is_const_def balanced_prob0_jozsa_algo prob0_jozsa_algo_of_balanced by auto end \ No newline at end of file diff --git a/thys/Isabelle_Marries_Dirac/More_Tensor.thy b/thys/Isabelle_Marries_Dirac/More_Tensor.thy --- a/thys/Isabelle_Marries_Dirac/More_Tensor.thy +++ b/thys/Isabelle_Marries_Dirac/More_Tensor.thy @@ -1,518 +1,518 @@ (* Authors: Anthony Bordg, University of Cambridge, apdb3@cam.ac.uk; Yijun He, University of Cambridge, yh403@cam.ac.uk *) section \Further Results on Tensor Products\ theory More_Tensor imports Quantum Tensor Jordan_Normal_Form.Matrix Basics begin lemma tensor_prod_2 [simp]: "mult.vec_vec_Tensor (*) [x1::complex,x2] [x3, x4] = [x1 * x3, x1 * x4, x2 * x3, x2 * x4]" proof - have "Matrix_Tensor.mult (1::complex) (*)" by (simp add: Matrix_Tensor.mult_def) thus "mult.vec_vec_Tensor (*) [x1::complex,x2] [x3,x4] = [x1*x3,x1*x4,x2*x3,x2*x4]" using mult.vec_vec_Tensor_def[of "(1::complex)" "(*)"] mult.times_def[of "(1::complex)" "(*)"] by simp qed lemma list_vec [simp]: assumes "v \ state_qbit 1" shows "list_of_vec v = [v $ 0, v $ 1]" proof - have "Rep_vec v = (fst(Rep_vec v), snd(Rep_vec v))" by simp also have "\ = (2, vec_index v)" by (metis (mono_tags, lifting) assms dim_vec.rep_eq mem_Collect_eq power_one_right state_qbit_def vec_index.rep_eq) moreover have "[0..<2::nat] = [0,1]" by(simp add: upt_rec) ultimately show ?thesis by(simp add: list_of_vec_def) qed lemma vec_tensor_prod_2 [simp]: assumes "v \ state_qbit 1" and "w \ state_qbit 1" shows "v \ w = vec_of_list [v $ 0 * w $ 0, v $ 0 * w $ 1, v $ 1 * w $ 0, v $ 1 * w $ 1]" proof - have "list_of_vec v = [v $ 0, v $ 1]" using assms by simp moreover have "list_of_vec w = [w $ 0, w $ 1]" using assms by simp ultimately show "v \ w = vec_of_list [v $ 0 * w $ 0, v $ 0 * w $ 1, v $ 1 * w $ 0, v $ 1 * w $ 1]" by(simp add: tensor_vec_def) qed lemma vec_dim_of_vec_of_list [simp]: assumes "length l = n" shows "dim_vec (vec_of_list l) = n" using assms vec_of_list_def by simp lemma vec_tensor_prod_2_bis [simp]: assumes "v \ state_qbit 1" and "w \ state_qbit 1" shows "v \ w = Matrix.vec 4 (\i. if i = 0 then v $ 0 * w $ 0 else if i = 3 then v $ 1 * w $ 1 else if i = 1 then v $ 0 * w $ 1 else v $ 1 * w $ 0)" proof define u where "u = Matrix.vec 4 (\i. if i = 0 then v $ 0 * w $ 0 else if i = 3 then v $ 1 * w $ 1 else if i = 1 then v $ 0 * w $ 1 else v $ 1 * w $ 0)" then show f2:"dim_vec (v \ w) = dim_vec u" using assms by simp show "\i. i < dim_vec u \ (v \ w) $ i = u $ i" apply (auto simp: u_def) using assms apply auto[3] apply (simp add: numeral_3_eq_3) using u_def vec_of_list_index vec_tensor_prod_2 index_is_2 by (metis (no_types, lifting) One_nat_def assms nth_Cons_0 nth_Cons_Suc numeral_2_eq_2) qed lemma index_col_mat_of_cols_list [simp]: assumes "i < n" and "j < length l" shows "Matrix.col (mat_of_cols_list n l) j $ i = l ! j ! i" apply (auto simp: Matrix.col_def mat_of_cols_list_def) using assms less_le_trans by fastforce lemma multTensor2 [simp]: assumes a1:"A = Matrix.mat 2 1 (\(i,j). if i = 0 then a0 else a1)" and a2:"B = Matrix.mat 2 1 (\(i,j). if i = 0 then b0 else b1)" shows "mult.Tensor (*) (mat_to_cols_list A) (mat_to_cols_list B) = [[a0*b0, a0*b1, a1*b0, a1*b1]]" proof - have "mat_to_cols_list A = [[a0, a1]]" by (auto simp: a1 mat_to_cols_list_def) (simp add: numeral_2_eq_2) moreover have f2:"mat_to_cols_list B = [[b0, b1]]" by (auto simp: a2 mat_to_cols_list_def) (simp add: numeral_2_eq_2) ultimately have "mult.Tensor (*) (mat_to_cols_list A) (mat_to_cols_list B) = mult.Tensor (*) [[a0,a1]] [[b0,b1]]" by simp thus ?thesis using mult.Tensor_def[of "(1::complex)" "(*)"] mult.times_def[of "(1::complex)" "(*)"] by (metis (mono_tags, lifting) append_self_conv list.simps(6) mult.Tensor.simps(2) mult.vec_mat_Tensor.simps(1) mult.vec_mat_Tensor.simps(2) plus_mult_cpx plus_mult_def tensor_prod_2) qed lemma multTensor2_bis [simp]: assumes a1:"dim_row A = 2" and a2:"dim_col A = 1" and a3:"dim_row B = 2" and a4:"dim_col B = 1" shows "mult.Tensor (*) (mat_to_cols_list A) (mat_to_cols_list B) = [[A $$ (0,0) * B $$ (0,0), A $$ (0,0) * B $$ (1,0), A $$ (1,0) * B $$ (0,0), A $$ (1,0) * B $$ (1,0)]]" proof - have "mat_to_cols_list A = [[A $$ (0,0), A $$ (1,0)]]" by (auto simp: a1 a2 mat_to_cols_list_def) (simp add: numeral_2_eq_2) moreover have f2:"mat_to_cols_list B = [[B $$ (0,0), B $$ (1,0)]]" by (auto simp: a3 a4 mat_to_cols_list_def) (simp add: numeral_2_eq_2) ultimately have "mult.Tensor (*) (mat_to_cols_list A) (mat_to_cols_list B) = mult.Tensor (*) [[A $$ (0,0), A $$ (1,0)]] [[B $$ (0,0), B $$ (1,0)]]" by simp thus ?thesis using mult.Tensor_def[of "(1::complex)" "(*)"] mult.times_def[of "(1::complex)" "(*)"] - by (smt append_self_conv list.simps(6) mult.Tensor.simps(2) mult.vec_mat_Tensor.simps(1) + by (smt (verit) append_self_conv list.simps(6) mult.Tensor.simps(2) mult.vec_mat_Tensor.simps(1) mult.vec_mat_Tensor.simps(2) plus_mult_cpx plus_mult_def tensor_prod_2) qed lemma mat_tensor_prod_2_prelim [simp]: assumes "state 1 v" and "state 1 w" shows "v \ w = mat_of_cols_list 4 [[v $$ (0,0) * w $$ (0,0), v $$ (0,0) * w $$ (1,0), v $$ (1,0) * w $$ (0,0), v $$ (1,0) * w $$ (1,0)]]" proof define u where "u = mat_of_cols_list 4 [[v $$ (0,0) * w $$ (0,0), v $$ (0,0) * w $$ (1,0), v $$ (1,0) * w $$ (0,0), v $$ (1,0) * w $$ (1,0)]]" then show f1:"dim_row (v \ w) = dim_row u" using assms state_def mat_of_cols_list_def tensor_mat_def by simp show f2:"dim_col (v \ w) = dim_col u" using assms state_def mat_of_cols_list_def tensor_mat_def u_def by simp show "\i j. i < dim_row u \ j < dim_col u \ (v \ w) $$ (i, j) = u $$ (i, j)" using u_def tensor_mat_def assms state_def by simp qed lemma mat_tensor_prod_2_col [simp]: assumes "state 1 v" and "state 1 w" shows "Matrix.col (v \ w) 0 = Matrix.col v 0 \ Matrix.col w 0" proof show f1:"dim_vec (Matrix.col (v \ w) 0) = dim_vec (Matrix.col v 0 \ Matrix.col w 0)" using assms vec_tensor_prod_2_bis - by (smt Tensor.mat_of_cols_list_def dim_col dim_row_mat(1) dim_vec mat_tensor_prod_2_prelim state.state_to_state_qbit) + by (smt (verit) Tensor.mat_of_cols_list_def dim_col dim_row_mat(1) dim_vec mat_tensor_prod_2_prelim state.state_to_state_qbit) next show "\i. i Matrix.col w 0) \ Matrix.col (v \ w) 0 $ i = (Matrix.col v 0 \ Matrix.col w 0) $ i" proof - have "dim_vec (Matrix.col v 0 \ Matrix.col w 0) = 4" by (metis (no_types, lifting) assms(1) assms(2) dim_vec state.state_to_state_qbit vec_tensor_prod_2_bis) moreover have "(Matrix.col v 0 \ Matrix.col w 0) $ 0 = v $$ (0,0) * w $$ (0,0)" using assms vec_tensor_prod_2 state.state_to_state_qbit col_index_of_mat_col - by (smt nth_Cons_0 power_one_right state_def vec_of_list_index zero_less_numeral) + by (smt (verit) nth_Cons_0 power_one_right state_def vec_of_list_index zero_less_numeral) moreover have "\ = Matrix.col (v \ w) 0 $ 0" using assms by simp moreover have "(Matrix.col v 0 \ Matrix.col w 0) $ 1 = v $$ (0,0) * w $$ (1,0)" using assms vec_tensor_prod_2 state.state_to_state_qbit col_index_of_mat_col - by (smt One_nat_def Suc_1 lessI nth_Cons_0 power_one_right state_def vec_index_vCons_Suc + by (smt (verit) One_nat_def Suc_1 lessI nth_Cons_0 power_one_right state_def vec_index_vCons_Suc vec_of_list_Cons vec_of_list_index zero_less_numeral) moreover have "\ = Matrix.col (v \ w) 0 $ 1" using assms by simp moreover have "(Matrix.col v 0 \ Matrix.col w 0) $ 2 = v $$ (1,0) * w $$ (0,0)" using assms vec_tensor_prod_2 state.state_to_state_qbit col_index_of_mat_col - by (smt One_nat_def Suc_1 lessI nth_Cons_0 power_one_right state_def vec_index_vCons_Suc + by (smt (verit) One_nat_def Suc_1 lessI nth_Cons_0 power_one_right state_def vec_index_vCons_Suc vec_of_list_Cons vec_of_list_index zero_less_numeral) moreover have "\ = Matrix.col (v \ w) 0 $ 2" using assms by simp moreover have "(Matrix.col v 0 \ Matrix.col w 0) $ 3 = v $$ (1,0) * w $$ (1,0)" using assms vec_tensor_prod_2 state.state_to_state_qbit col_index_of_mat_col numeral_3_eq_3 - by (smt One_nat_def Suc_1 lessI nth_Cons_0 power_one_right state_def vec_index_vCons_Suc + by (smt (verit) One_nat_def Suc_1 lessI nth_Cons_0 power_one_right state_def vec_index_vCons_Suc vec_of_list_Cons vec_of_list_index zero_less_numeral) moreover have "\ = Matrix.col (v \ w) 0 $ 3" using assms by simp ultimately show "\i. i Matrix.col w 0) \ Matrix.col (v \ w) 0 $ i = (Matrix.col v 0 \ Matrix.col w 0) $ i" using index_sl_four by auto qed qed lemma mat_tensor_prod_2 [simp]: assumes "state 1 v" and "state 1 w" shows "v \ w = Matrix.mat 4 1 (\(i,j). if i = 0 then v $$ (0,0) * w $$ (0,0) else if i = 3 then v $$ (1,0) * w $$ (1,0) else if i = 1 then v $$ (0,0) * w $$ (1,0) else v $$ (1,0) * w $$ (0,0))" proof define u where "u = Matrix.mat 4 1 (\(i,j). if i = 0 then v $$ (0,0) * w $$ (0,0) else if i = 3 then v $$ (1,0) * w $$ (1,0) else if i = 1 then v $$ (0,0) * w $$ (1,0) else v $$ (1,0) * w $$ (0,0))" then show "dim_row (v \ w) = dim_row u" using assms tensor_mat_def state_def by(simp add: Tensor.mat_of_cols_list_def) also have "\ = 4" by (simp add: u_def) show "dim_col (v \ w) = dim_col u" using u_def assms tensor_mat_def state_def Tensor.mat_of_cols_list_def by simp moreover have "\ = 1" by(simp add: u_def) ultimately show "\i j. i < dim_row u \ j < dim_col u \ (v \ w) $$ (i, j) = u $$ (i,j)" proof - fix i j::nat assume a1:"i < dim_row u" and a2:"j < dim_col u" then have "(v \ w) $$ (i, j) = Matrix.col (v \ w) 0 $ i" using Matrix.col_def u_def assms by simp then have f1:"(v \ w) $$ (i, j) = (Matrix.col v 0 \ Matrix.col w 0) $ i" using assms mat_tensor_prod_2_col by simp have "(Matrix.col v 0 \ Matrix.col w 0) $ i = Matrix.vec 4 (\i. if i = 0 then Matrix.col v 0 $ 0 * Matrix.col w 0 $ 0 else if i = 3 then Matrix.col v 0 $ 1 * Matrix.col w 0 $ 1 else if i = 1 then Matrix.col v 0 $ 0 * Matrix.col w 0 $ 1 else Matrix.col v 0 $ 1 * Matrix.col w 0 $ 0) $ i" using vec_tensor_prod_2_bis assms state.state_to_state_qbit by presburger thus "(v \ w) $$ (i, j) = u $$ (i,j)" using u_def a1 a2 assms state_def by simp qed qed lemma mat_tensor_prod_2_bis: assumes "state 1 v" and "state 1 w" shows "v \ w = |Matrix.vec 4 (\i. if i = 0 then v $$ (0,0) * w $$ (0,0) else if i = 3 then v $$ (1,0) * w $$ (1,0) else if i = 1 then v $$ (0,0) * w $$ (1,0) else v $$ (1,0) * w $$ (0,0))\" using assms ket_vec_def mat_tensor_prod_2 by(simp add: mat_eq_iff) lemma eq_ket_vec: fixes u v:: "complex Matrix.vec" assumes "u = v" shows "|u\ = |v\" using assms by simp lemma mat_tensor_ket_vec: assumes "state 1 v" and "state 1 w" shows "v \ w = |(Matrix.col v 0) \ (Matrix.col w 0)\" proof - have "v \ w = |Matrix.col v 0\ \ |Matrix.col w 0\" using assms state_def by simp also have "\ = |Matrix.vec 4 (\i. if i = 0 then |Matrix.col v 0\ $$ (0,0) * |Matrix.col w 0\ $$ (0,0) else if i = 3 then |Matrix.col v 0\ $$ (1,0) * |Matrix.col w 0\ $$ (1,0) else if i = 1 then |Matrix.col v 0\ $$ (0,0) * |Matrix.col w 0\ $$ (1,0) else |Matrix.col v 0\ $$ (1,0) * |Matrix.col w 0\ $$ (0,0))\" using assms mat_tensor_prod_2_bis state_col_ket_vec by simp also have "\ = |Matrix.vec 4 (\i. if i = 0 then v $$ (0,0) * w $$ (0,0) else if i = 3 then v $$ (1,0) * w $$ (1,0) else if i = 1 then v $$ (0,0) * w $$ (1,0) else v $$ (1,0) * w $$ (0,0))\" using assms mat_tensor_prod_2_bis calculation by auto also have "\ = |Matrix.vec 4 (\i. if i = 0 then Matrix.col v 0 $ 0 * Matrix.col w 0 $ 0 else if i = 3 then Matrix.col v 0 $ 1 * Matrix.col w 0 $ 1 else if i = 1 then Matrix.col v 0 $ 0 * Matrix.col w 0 $ 1 else Matrix.col v 0 $ 1 * Matrix.col w 0 $ 0)\" apply(rule eq_ket_vec) apply (rule eq_vecI) using col_index_of_mat_col assms state_def by auto finally show ?thesis using vec_tensor_prod_2_bis assms state.state_to_state_qbit by presburger qed text \The property of being a state (resp. a gate) is preserved by tensor product.\ lemma tensor_state2 [simp]: assumes "state 1 u" and "state 1 v" shows "state 2 (u \ v)" proof show "dim_col (u \ v) = 1" using assms dim_col_tensor_mat state.is_column by presburger show "dim_row (u \ v) = 2\<^sup>2" using assms dim_row_tensor_mat state.dim_row by (metis (mono_tags, lifting) power2_eq_square power_one_right) show "\Matrix.col (u \ v) 0\ = 1" proof - define l where d0:"l = [u $$ (0,0) * v $$ (0,0), u $$ (0,0) * v $$ (1,0), u $$ (1,0) * v $$ (0,0), u $$ (1,0) * v $$ (1,0)]" then have f4:"length l = 4" by simp also have "u \ v = mat_of_cols_list 4 [[u $$ (0,0) * v $$ (0,0), u $$ (0,0) * v $$ (1,0), u $$ (1,0) * v $$ (0,0), u $$ (1,0) * v $$ (1,0)]]" using assms by simp then have "Matrix.col (u \ v) 0 = vec_of_list [u $$ (0,0) * v $$ (0,0), u $$ (0,0) * v $$ (1,0), u $$ (1,0) * v $$ (0,0), u $$ (1,0) * v $$ (1,0)]" by (metis One_nat_def Suc_eq_plus1 add_Suc col_mat_of_cols_list list.size(3) list.size(4) nth_Cons_0 numeral_2_eq_2 numeral_Bit0 plus_1_eq_Suc vec_of_list_Cons zero_less_one_class.zero_less_one) then have f5:"\Matrix.col (u \ v) 0\ = sqrt(\i<4. (cmod (l ! i))\<^sup>2)" by (metis d0 f4 One_nat_def cpx_length_of_vec_of_list d0 vec_of_list_Cons) also have "\ = sqrt ((cmod (u $$ (0,0) * v $$ (0,0)))\<^sup>2 + (cmod(u $$ (0,0) * v $$ (1,0)))\<^sup>2 + (cmod(u $$ (1,0) * v $$ (0,0)))\<^sup>2 + (cmod(u $$ (1,0) * v $$ (1,0)))\<^sup>2)" proof - have "(\i<4. (cmod (l ! i))\<^sup>2) = (cmod (l ! 0))\<^sup>2 + (cmod (l ! 1))\<^sup>2 + (cmod (l ! 2))\<^sup>2 + (cmod (l ! 3))\<^sup>2" using sum_insert - by (smt One_nat_def empty_iff finite.emptyI finite.insertI insertE lessThan_0 lessThan_Suc + by (smt (verit) One_nat_def empty_iff finite.emptyI finite.insertI insertE lessThan_0 lessThan_Suc numeral_2_eq_2 numeral_3_eq_3 numeral_plus_one one_plus_numeral_commute plus_1_eq_Suc semiring_norm(2) semiring_norm(8) sum.empty) also have "\ = (cmod (u $$ (0,0) * v $$ (0,0)))\<^sup>2 + (cmod(u $$ (0,0) * v $$ (1,0)))\<^sup>2 + (cmod(u $$ (1,0) * v $$ (0,0)))\<^sup>2 + (cmod(u $$ (1,0) * v $$ (1,0)))\<^sup>2" using d0 by simp finally show ?thesis by(simp add: f5) qed moreover have "\ = sqrt ((cmod (u $$ (0,0)))\<^sup>2 * (cmod (v $$ (0,0)))\<^sup>2 + (cmod(u $$ (0,0)))\<^sup>2 * (cmod (v $$ (1,0)))\<^sup>2 + (cmod(u $$ (1,0)))\<^sup>2 * (cmod (v $$ (0,0)))\<^sup>2 + (cmod(u $$ (1,0)))\<^sup>2 * (cmod(v $$ (1,0)))\<^sup>2)" by (simp add: norm_mult power_mult_distrib) moreover have "\ = sqrt ((((cmod(u $$ (0,0)))\<^sup>2 + (cmod(u $$ (1,0)))\<^sup>2)) * (((cmod(v $$ (0,0)))\<^sup>2 + (cmod(v $$ (1,0)))\<^sup>2)))" by (simp add: distrib_left mult.commute) ultimately have f6:"\Matrix.col (u \ v) 0\\<^sup>2 = (((cmod(u $$ (0,0)))\<^sup>2 + (cmod(u $$ (1,0)))\<^sup>2)) * (((cmod(v $$ (0,0)))\<^sup>2 + (cmod(v $$ (1,0)))\<^sup>2))" by (simp add: f4) also have f7:"\ = (\i< 2. (cmod (u $$ (i,0)))\<^sup>2) * (\i<2. (cmod (v $$ (i,0)))\<^sup>2)" by (simp add: numeral_2_eq_2) also have f8:"\ = (\i< 2.(cmod (Matrix.col u 0 $ i))\<^sup>2) * (\i<2.(cmod (Matrix.col v 0 $ i))\<^sup>2)" using assms index_col state_def by simp finally show ?thesis proof - have f1:"(\i< 2.(cmod (Matrix.col u 0 $ i))\<^sup>2) = 1" using assms(1) state_def cpx_vec_length_def by auto have f2:"(\i< 2.(cmod (Matrix.col v 0 $ i))\<^sup>2) = 1" using assms(2) state_def cpx_vec_length_def by auto thus ?thesis using f1 f8 f5 f6 f7 by (simp add: \sqrt (\i<4. (cmod (l ! i))\<^sup>2) = sqrt ((cmod (u $$ (0, 0) * v $$ (0, 0)))\<^sup>2 + (cmod (u $$ (0, 0) * v $$ (1, 0)))\<^sup>2 + (cmod (u $$ (1, 0) * v $$ (0, 0)))\<^sup>2 + (cmod (u $$ (1, 0) * v $$ (1, 0)))\<^sup>2)\) qed qed qed lemma sum_prod: fixes f::"nat \ complex" and g::"nat \ complex" shows "(\iiji<(a+1)*b. f (i div b) * g (i mod b)) = (\ii\{a*b..<(a+1)*b}. f (i div b) * g (i mod b))" apply (auto simp: algebra_simps) - by (smt add.commute mult_Suc sum.lessThan_Suc sum.nat_group) + by (smt (verit) add.commute mult_Suc sum.lessThan_Suc sum.nat_group) also have "\ = (\iji\{a*b..<(a+1)*b}. f (i div b) * g (i mod b))" by (simp add: Suc.IH) also have "\ = (\iji\{a*b..<(a+1)*b}. f (a) * g(i-a*b))" by simp also have "\ = (\iji\{a*b..<(a+1)*b}. g(i-a*b))" by(simp add: sum_distrib_left) also have "\ = (\iji\{.. v)" proof show c1:"dim_col (u \ v) = 1" using assms dim_col_tensor_mat state.is_column by presburger show c2:"dim_row (u \ v) = 2^(m + n)" using assms dim_row_tensor_mat state.dim_row by (metis power_add) have "(\i<2^(m + n). (cmod (u $$ (i div 2 ^ n, 0) * v $$ (i mod 2 ^ n, 0)))\<^sup>2) = (\i<2^(m + n). (cmod (u $$ (i div 2 ^ n, 0)))\<^sup>2 * (cmod (v $$ (i mod 2 ^ n, 0)))\<^sup>2)" by (simp add: sqr_of_cmod_of_prod) also have "\ = (\i<2^m. (cmod (u $$ (i, 0)))\<^sup>2) * (\i<2^n. (cmod (v $$ (i, 0)))\<^sup>2)" proof- have "\ = (\i<2^(m + n).complex_of_real((cmod (u $$ (i div 2^n,0)))\<^sup>2) * complex_of_real((cmod (v $$ (i mod 2^n,0)))\<^sup>2))" by simp moreover have "(\i<2^m. (cmod (u $$ (i, 0)))\<^sup>2) = (\i<2^m. complex_of_real ((cmod (u $$ (i,0)))\<^sup>2))" by simp moreover have "(\i<2^n. (cmod (v $$ (i, 0)))\<^sup>2) = (\i<2^n. complex_of_real ((cmod (v $$ (i, 0)))\<^sup>2))" by simp ultimately show ?thesis using sum_prod[of "\i. (cmod (u $$ (i , 0)))\<^sup>2" "2^n" "\i. (cmod (v $$ (i , 0)))\<^sup>2" "2^m"] - by (smt of_real_eq_iff of_real_mult power_add) + by (smt (verit) of_real_eq_iff of_real_mult power_add) qed also have "\ = 1" proof- have "(\i<2^m. (cmod (u $$ (i, 0)))\<^sup>2) = 1" using assms(1) state_def cpx_vec_length_def by auto moreover have "(\i<2^n. (cmod (v $$ (i, 0)))\<^sup>2) = 1" using assms(2) state_def cpx_vec_length_def by auto ultimately show ?thesis by simp qed ultimately show "\Matrix.col (u \ v) 0\ = 1" using c1 c2 assms state_def by (auto simp add: cpx_vec_length_def) qed lemma dim_row_of_tensor_gate: assumes "gate m G1" and "gate n G2" shows "dim_row (G1 \ G2) = 2^(m+n)" using assms dim_row_tensor_mat gate.dim_row by (simp add: power_add) lemma tensor_gate_sqr_mat: assumes "gate m G1" and "gate n G2" shows "square_mat (G1 \ G2)" using assms gate.square_mat by simp lemma dim_row_of_one_mat_less_pow: assumes "gate m G1" and "gate n G2" and "i < dim_row (1\<^sub>m(dim_col G1 * dim_col G2))" shows "i < 2^(m+n)" using assms gate_def by (simp add: power_add) lemma dim_col_of_one_mat_less_pow: assumes "gate m G1" and "gate n G2" and "j < dim_col (1\<^sub>m(dim_col G1 * dim_col G2))" shows "j < 2^(m+n)" using assms gate_def by (simp add: power_add) lemma index_tensor_gate_unitary1: assumes "gate m G1" and "gate n G2" and "i < dim_row (1\<^sub>m(dim_col G1 * dim_col G2))" and "j < dim_col (1\<^sub>m(dim_col G1 * dim_col G2))" shows "((G1 \ G2)\<^sup>\ * (G1 \ G2)) $$ (i, j) = 1\<^sub>m(dim_col G1 * dim_col G2) $$ (i, j)" proof- have "\k. k<2^(m+n) \ cnj((G1 \ G2) $$ (k,i)) = cnj(G1 $$ (k div 2^n, i div 2^n)) * cnj(G2 $$ (k mod 2^n, i mod 2^n))" using assms(1-3) by (simp add: gate_def power_add) moreover have "\k. k<2^(m+n) \ (G1 \ G2) $$ (k,j) = G1 $$ (k div 2^n, j div 2^n) * G2 $$ (k mod 2^n, j mod 2^n)" using assms(1,2,4) by (simp add: gate_def power_add) ultimately have "\k. k<2^(m+n) \ cnj((G1 \ G2) $$ (k,i)) * ((G1 \ G2) $$ (k,j)) = cnj(G1 $$ (k div 2^n, i div 2^n)) * G1 $$ (k div 2^n, j div 2^n) * cnj(G2 $$ (k mod 2^n, i mod 2^n)) * G2 $$ (k mod 2^n, j mod 2^n)" by simp then have "(\k<2^(m+n). cnj((G1 \ G2) $$ (k,i)) * ((G1 \ G2) $$ (k,j))) = (\k<2^(m+n). cnj(G1 $$ (k div 2^n, i div 2^n)) * G1 $$ (k div 2^n, j div 2^n) * cnj(G2 $$ (k mod 2^n, i mod 2^n)) * G2 $$ (k mod 2^n, j mod 2^n))" by simp also have "\ = (\k<2^m. cnj(G1 $$ (k, i div 2^n)) * G1 $$ (k, j div 2^n)) * (\k<2^n. cnj(G2 $$ (k, i mod 2^n)) * G2 $$ (k, j mod 2^n))" using sum_prod[of "\x. cnj(G1 $$ (x, i div 2^n)) * G1 $$ (x, j div 2^n)" "2^n" "\x. cnj(G2 $$ (x, i mod 2^n)) * G2 $$ (x, j mod 2^n)" "2^m"] by (metis (no_types, lifting) power_add semigroup_mult_class.mult.assoc sum.cong) also have "((G1 \ G2)\<^sup>\ * (G1 \ G2)) $$ (i, j) = (\k<2^(m+n).cnj((G1 \ G2) $$ (k,i)) * ((G1 \ G2) $$ (k,j)))" using assms index_matrix_prod[of "i" "(G1 \ G2)\<^sup>\" "j" "(G1 \ G2)"] dagger_def dim_row_of_tensor_gate tensor_gate_sqr_mat by simp ultimately have "((G1 \ G2)\<^sup>\ * (G1 \ G2)) $$ (i,j) = (\k1<2^m. cnj(G1 $$ (k1, i div 2^n)) * G1 $$ (k1, j div 2^n)) * (\k2<2^n. cnj(G2 $$ (k2, i mod 2^n)) * G2 $$ (k2, j mod 2^n))" by simp moreover have "(\k<2^m. cnj(G1 $$ (k, i div 2^n))* G1 $$ (k, j div 2^n)) = (G1\<^sup>\ * G1) $$ (i div 2^n, j div 2^n)" using assms gate_def dagger_def index_matrix_prod[of "i div 2^n" "G1\<^sup>\" "j div 2^n" "G1"] by (simp add: less_mult_imp_div_less power_add) moreover have "\ = 1\<^sub>m(2^m) $$ (i div 2^n, j div 2^n)" using assms(1,2) gate_def unitary_def by simp moreover have "(\k<2^n. cnj(G2 $$ (k, i mod 2^n))* G2 $$ (k, j mod 2^n)) = (G2\<^sup>\ * G2) $$ (i mod 2^n, j mod 2^n)" using assms(1,2) gate_def dagger_def index_matrix_prod[of "i mod 2^n" "G2\<^sup>\" "j mod 2^n" "G2"] by simp moreover have "\ = 1\<^sub>m(2^n) $$ (i mod 2^n, j mod 2^n)" using assms(1,2) gate_def unitary_def by simp ultimately have "((G1 \ G2)\<^sup>\ * (G1 \ G2)) $$ (i,j) = 1\<^sub>m (2^m) $$ (i div 2^n, j div 2^n) * 1\<^sub>m (2^n) $$ (i mod 2^n, j mod 2^n)" by simp thus ?thesis using assms assms(3,4) gate_def index_one_mat_div_mod[of "i" "m" "n" "j"] by(simp add: power_add) qed lemma tensor_gate_unitary1 [simp]: assumes "gate m G1" and "gate n G2" shows "(G1 \ G2)\<^sup>\ * (G1 \ G2) = 1\<^sub>m(dim_col G1 * dim_col G2)" proof show "dim_row ((G1 \ G2)\<^sup>\ * (G1 \ G2)) = dim_row (1\<^sub>m(dim_col G1 * dim_col G2))" by simp show "dim_col ((G1 \ G2)\<^sup>\ * (G1 \ G2)) = dim_col (1\<^sub>m(dim_col G1 * dim_col G2))" by simp fix i j assume "i < dim_row (1\<^sub>m(dim_col G1 * dim_col G2))" and "j < dim_col (1\<^sub>m(dim_col G1 * dim_col G2))" thus "((G1 \ G2)\<^sup>\ * (G1 \ G2)) $$ (i, j) = 1\<^sub>m(dim_col G1 * dim_col G2) $$ (i, j)" using assms index_tensor_gate_unitary1 by simp qed lemma index_tensor_gate_unitary2 [simp]: assumes "gate m G1" and "gate n G2" and "i < dim_row (1\<^sub>m (dim_col G1 * dim_col G2))" and "j < dim_col (1\<^sub>m (dim_col G1 * dim_col G2))" shows "((G1 \ G2) * ((G1 \ G2)\<^sup>\)) $$ (i, j) = 1\<^sub>m(dim_col G1 * dim_col G2) $$ (i, j)" proof- have "\k. k<2^(m+n) \ (G1 \ G2) $$ (i,k) = G1 $$ (i div 2^n, k div 2^n) * G2 $$ (i mod 2^n, k mod 2^n)" using assms(1-3) by (simp add: gate_def power_add) moreover have "\k. k<2^(m+n) \ cnj((G1 \ G2) $$ (j,k)) = cnj(G1 $$ (j div 2^n, k div 2^n)) * cnj(G2 $$ (j mod 2^n, k mod 2^n))" using assms(1,2,4) by (simp add: gate_def power_add) ultimately have "\k. k\{..<2^(m+n)} \ (G1 \ G2) $$ (i,k) * cnj((G1 \ G2) $$ (j,k)) = G1 $$ (i div 2^n, k div 2^n) * cnj(G1 $$ (j div 2^n, k div 2^n)) * G2 $$ (i mod 2^n, k mod 2^n) * cnj(G2 $$ (j mod 2^n, k mod 2^n))" by simp then have "(\k<2^(m+n). (G1 \ G2) $$ (i,k) * cnj((G1 \ G2) $$ (j,k))) = (\k<2^(m+n). G1 $$ (i div 2^n, k div 2^n) * cnj(G1 $$ (j div 2^n, k div 2^n)) * G2 $$ (i mod 2^n, k mod 2^n) * cnj(G2 $$ (j mod 2^n, k mod 2^n)))" by simp also have "\ = (\k<2^m. G1 $$ (i div 2^n, k) * cnj(G1 $$ (j div 2^n, k))) * (\k<2^n. G2 $$ (i mod 2^n, k) * cnj(G2 $$ (j mod 2^n, k)))" using sum_prod[of "\k. G1 $$ (i div 2^n, k) * cnj(G1 $$ (j div 2^n, k))" "2^n" "\k. G2 $$ (i mod 2^n, k) * cnj(G2 $$ (j mod 2^n, k))" "2^m"] by (metis (no_types, lifting) power_add semigroup_mult_class.mult.assoc sum.cong) also have "((G1 \ G2) * ((G1 \ G2)\<^sup>\)) $$ (i, j) = (\k<2^(m+n). (G1 \ G2) $$ (i,k) * cnj((G1 \ G2) $$ (j,k)))" using assms index_matrix_prod[of "i" "(G1 \ G2)" "j" "(G1 \ G2)\<^sup>\"] dagger_def dim_row_of_tensor_gate tensor_gate_sqr_mat by simp ultimately have "((G1 \ G2) * ((G1 \ G2)\<^sup>\)) $$ (i,j) = (\k<2^m. G1 $$ (i div 2^n, k) * cnj(G1 $$ (j div 2^n, k))) * (\k<2^n. G2 $$ (i mod 2^n, k) * cnj(G2 $$ (j mod 2^n, k)))" by simp moreover have "(\k<2^m. G1 $$ (i div 2^n, k) * cnj(G1 $$ (j div 2^n, k))) = (G1 * (G1\<^sup>\)) $$ (i div 2^n, j div 2^n)" using assms gate_def dagger_def index_matrix_prod[of "i div 2^n" "G1" "j div 2^n" "G1\<^sup>\"] by (simp add: less_mult_imp_div_less power_add) moreover have "\ = 1\<^sub>m(2^m) $$ (i div 2^n, j div 2^n)" using assms(1,2) gate_def unitary_def by simp moreover have "(\k<2^n. G2 $$ (i mod 2^n, k) * cnj(G2 $$ (j mod 2^n, k))) = (G2 * (G2\<^sup>\)) $$ (i mod 2^n, j mod 2^n)" using assms(1,2) gate_def dagger_def index_matrix_prod[of "i mod 2^n" "G2" "j mod 2^n" "G2\<^sup>\"] by simp moreover have "\ = 1\<^sub>m(2^n) $$ (i mod 2^n, j mod 2^n)" using assms(1,2) gate_def unitary_def by simp ultimately have "((G1 \ G2) * ((G1 \ G2)\<^sup>\)) $$ (i,j) = 1\<^sub>m(2^m) $$ (i div 2^n, j div 2^n) * 1\<^sub>m(2^n) $$ (i mod 2^n, j mod 2^n)" by simp thus ?thesis using assms gate_def index_one_mat_div_mod[of "i" "m" "n" "j"] by(simp add: power_add) qed lemma tensor_gate_unitary2 [simp]: assumes "gate m G1" and "gate n G2" shows "(G1 \ G2) * ((G1 \ G2)\<^sup>\) = 1\<^sub>m(dim_col G1 * dim_col G2)" proof show "dim_row ((G1 \ G2) * ((G1 \ G2)\<^sup>\)) = dim_row(1\<^sub>m (dim_col G1 * dim_col G2))" using assms gate_def by simp show "dim_col ((G1 \ G2) * ((G1 \ G2)\<^sup>\)) = dim_col (1\<^sub>m(dim_col G1 * dim_col G2))" using assms gate_def by simp fix i j assume "i < dim_row (1\<^sub>m (dim_col G1 * dim_col G2))" and "j < dim_col (1\<^sub>m (dim_col G1 * dim_col G2))" thus "((G1 \ G2) * ((G1 \ G2)\<^sup>\)) $$ (i, j) = 1\<^sub>m(dim_col G1 * dim_col G2) $$ (i, j)" using assms index_tensor_gate_unitary2 by simp qed lemma tensor_gate [simp]: assumes "gate m G1" and "gate n G2" shows "gate (m + n) (G1 \ G2)" proof show "dim_row (G1 \ G2) = 2^(m+n)" using assms dim_row_tensor_mat gate.dim_row by (simp add: power_add) show "square_mat (G1 \ G2)" using assms gate.square_mat by simp thus "unitary (G1 \ G2)" using assms unitary_def by simp qed end \ No newline at end of file diff --git a/thys/Isabelle_Marries_Dirac/No_Cloning.thy b/thys/Isabelle_Marries_Dirac/No_Cloning.thy --- a/thys/Isabelle_Marries_Dirac/No_Cloning.thy +++ b/thys/Isabelle_Marries_Dirac/No_Cloning.thy @@ -1,242 +1,242 @@ (* Authors: Anthony Bordg, University of Cambridge, apdb3@cam.ac.uk; Yijun He, University of Cambridge, yh403@cam.ac.uk *) section \The No-Cloning Theorem\ theory No_Cloning imports Quantum Tensor begin subsection \The Cauchy-Schwarz Inequality\ lemma inner_prod_expand: assumes "dim_vec a = dim_vec b" and "dim_vec a = dim_vec c" and "dim_vec a = dim_vec d" shows "\a + b|c + d\ = \a|c\ + \a|d\ + \b|c\ + \b|d\" apply (simp add: inner_prod_def) using assms sum.cong by (simp add: sum.distrib algebra_simps) lemma inner_prod_distrib_left: assumes "dim_vec a = dim_vec b" shows "\c \\<^sub>v a|b\ = cnj(c) * \a|b\" using assms inner_prod_def by (simp add: algebra_simps mult_hom.hom_sum) lemma inner_prod_distrib_right: assumes "dim_vec a = dim_vec b" shows "\a|c \\<^sub>v b\ = c * \a|b\" using assms by (simp add: algebra_simps mult_hom.hom_sum) lemma cauchy_schwarz_ineq: assumes "dim_vec v = dim_vec w" shows "(cmod(\v|w\))\<^sup>2 \ Re (\v|v\ * \w|w\)" proof (cases "\v|v\ = 0") case c0:True then have "\i. i < dim_vec v \ v $ i = 0" by(metis index_zero_vec(1) inner_prod_with_itself_nonneg_reals_non0) then have "(cmod(\v|w\))\<^sup>2 = 0" by (simp add: assms inner_prod_def) moreover have "Re (\v|v\ * \w|w\) = 0" by (simp add: c0) ultimately show ?thesis by simp next case c1:False have "dim_vec w = dim_vec (- \v|w\ / \v|v\ \\<^sub>v v)" by (simp add: assms) then have "\w + -\v|w\/\v|v\ \\<^sub>v v|w + -\v|w\/\v|v\ \\<^sub>v v\ = \w|w\ + \w|-\v|w\/\v|v\ \\<^sub>v v\ + \-\v|w\/\v|v\ \\<^sub>v v|w\ + \-\v|w\/\v|v\ \\<^sub>v v|-\v|w\/\v|v\ \\<^sub>v v\" using inner_prod_expand[of "w" "-\v|w\/\v|v\ \\<^sub>v v" "w" "-\v|w\/\v|v\ \\<^sub>v v"] by auto moreover have "\w|-\v|w\/\v|v\ \\<^sub>v v\ = -\v|w\/\v|v\ * \w|v\" using assms inner_prod_distrib_right[of "w" "v" "-\v|w\/\v|v\"] by simp moreover have "\-\v|w\/\v|v\ \\<^sub>v v|w\ = cnj(-\v|w\/\v|v\) * \v|w\" using assms inner_prod_distrib_left[of "v" "w" "-\v|w\/\v|v\"] by simp moreover have "\-\v|w\/\v|v\ \\<^sub>v v|-\v|w\/\v|v\ \\<^sub>v v\ = cnj(-\v|w\/\v|v\) * (-\v|w\/\v|v\) * \v|v\" using inner_prod_distrib_left[of "v" "-\v|w\/\v|v\ \\<^sub>v v" "-\v|w\/\v|v\"] inner_prod_distrib_right[of "v" "v" "-\v|w\/\v|v\"] by simp ultimately have "\w + -\v|w\/\v|v\ \\<^sub>v v|w + -\v|w\/\v|v\ \\<^sub>v v\ = \w|w\ - cmod(\v|w\)^2 / \v|v\" using assms inner_prod_cnj[of "w" "v"] inner_prod_cnj[of "v" "v"] complex_norm_square by simp moreover have "Re(\w + -\v|w\/\v|v\ \\<^sub>v v|w + -\v|w\/\v|v\ \\<^sub>v v\) \ 0" using inner_prod_with_itself_Re by blast ultimately have "Re(\w|w\) \ cmod(\v|w\)^2/Re(\v|v\)" using inner_prod_with_itself_real by simp moreover have c2:"Re(\v|v\) > 0" using inner_prod_with_itself_Re_non0 inner_prod_with_itself_eq0 c1 by auto ultimately have "Re(\w|w\) * Re(\v|v\) \ cmod(\v|w\)^2/Re(\v|v\) * Re(\v|v\)" by (metis less_numeral_extra(3) nonzero_divide_eq_eq pos_divide_le_eq) thus ?thesis using inner_prod_with_itself_Im c2 by (simp add: mult.commute) qed lemma cauchy_schwarz_eq [simp]: assumes "v = (l \\<^sub>v w)" shows "(cmod(\v|w\))\<^sup>2 = Re (\v|v\ * \w|w\)" proof- have "cmod(\v|w\) = cmod(cnj(l) * \w|w\)" using assms inner_prod_distrib_left[of "w" "w" "l"] by simp then have "cmod(\v|w\)^2 = cmod(l)^2 * \w|w\ * \w|w\" using complex_norm_square inner_prod_cnj[of "w" "w"] by simp moreover have "\v|v\ = cmod(l)^2 * \w|w\" using assms complex_norm_square inner_prod_distrib_left[of "w" "v" "l"] inner_prod_distrib_right[of "w" "w" "l"] by simp ultimately show ?thesis by (metis Re_complex_of_real) qed lemma cauchy_schwarz_col [simp]: assumes "dim_vec v = dim_vec w" and "(cmod(\v|w\))\<^sup>2 = Re (\v|v\ * \w|w\)" shows "\l. v = (l \\<^sub>v w) \ w = (l \\<^sub>v v)" proof (cases "\v|v\ = 0") case c0:True then have "\i. i < dim_vec v \ v $ i = 0" by(metis index_zero_vec(1) inner_prod_with_itself_nonneg_reals_non0) then have "v = 0 \\<^sub>v w" by (auto simp: assms) then show ?thesis by auto next case c1:False have f0:"dim_vec w = dim_vec (- \v|w\ / \v|v\ \\<^sub>v v)" by (simp add: assms(1)) then have "\w + -\v|w\/\v|v\ \\<^sub>v v|w + -\v|w\/\v|v\ \\<^sub>v v\ = \w|w\ + \w|-\v|w\/\v|v\ \\<^sub>v v\ + \-\v|w\/\v|v\ \\<^sub>v v|w\ + \-\v|w\/\v|v\ \\<^sub>v v|-\v|w\/\v|v\ \\<^sub>v v\" using inner_prod_expand[of "w" "-\v|w\/\v|v\ \\<^sub>v v" "w" "-\v|w\/\v|v\ \\<^sub>v v"] by simp moreover have "\w|-\v|w\/\v|v\ \\<^sub>v v\ = -\v|w\/\v|v\ * \w|v\" using assms(1) inner_prod_distrib_right[of "w" "v" "-\v|w\/\v|v\"] by simp moreover have "\-\v|w\/\v|v\ \\<^sub>v v|w\ = cnj(-\v|w\/\v|v\) * \v|w\" using assms(1) inner_prod_distrib_left[of "v" "w" "-\v|w\/\v|v\"] by simp moreover have "\-\v|w\/\v|v\ \\<^sub>v v|-\v|w\/\v|v\ \\<^sub>v v\ = cnj(-\v|w\/\v|v\) * (-\v|w\/\v|v\) * \v|v\" using inner_prod_distrib_left[of "v" "-\v|w\/\v|v\ \\<^sub>v v" "-\v|w\/\v|v\"] inner_prod_distrib_right[of "v" "v" "-\v|w\/\v|v\"] by simp ultimately have "\w + -\v|w\/\v|v\ \\<^sub>v v|w + -\v|w\/\v|v\ \\<^sub>v v\ = \w|w\ - cmod(\v|w\)^2 / \v|v\" using inner_prod_cnj[of "w" "v"] inner_prod_cnj[of "v" "v"] assms(1) complex_norm_square by simp moreover have "\w|w\ = cmod(\v|w\)^2 / \v|v\" using assms(2) inner_prod_with_itself_real by(metis Reals_mult c1 nonzero_mult_div_cancel_left of_real_Re) ultimately have "\w + -\v|w\/\v|v\ \\<^sub>v v|w + -\v|w\/\v|v\ \\<^sub>v v\ = 0" by simp then have "\i. i (w + -\v|w\/\v|v\ \\<^sub>v v) $ i = 0" by (metis f0 index_add_vec(2) index_zero_vec(1) inner_prod_with_itself_nonneg_reals_non0) then have "\i. i w $ i + -\v|w\/\v|v\ * v $ i = 0" by (metis assms(1) f0 index_add_vec(1) index_smult_vec(1)) then have "\i. i w $ i = \v|w\/\v|v\ * v $ i" by simp then have "w = \v|w\/\v|v\ \\<^sub>v v" by (auto simp add: assms(1)) thus ?thesis by auto qed subsection \The No-Cloning Theorem\ lemma eq_from_inner_prod [simp]: assumes "dim_vec v = dim_vec w" and "\v|w\ = 1" and "\v|v\ = 1" and "\w|w\ = 1" shows "v = w" proof- have "(cmod(\v|w\))\<^sup>2 = Re (\v|v\ * \w|w\)" by (simp add: assms) then have f0:"\l. v = (l \\<^sub>v w) \ w = (l \\<^sub>v v)" by (simp add: assms(1)) then show ?thesis proof (cases "\l. v = (l \\<^sub>v w)") case True then have "\l. v = (l \\<^sub>v w) \ \v|w\ = cnj(l) * \w|w\" using inner_prod_distrib_left by auto then show ?thesis by (simp add: assms(2,4)) next case False then have "\l. w = (l \\<^sub>v v) \ \v|w\ = l * \v|v\" using f0 inner_prod_distrib_right by auto then show ?thesis by (simp add: assms(2,3)) qed qed lemma hermite_cnj_of_tensor: shows "(A \ B)\<^sup>\ = (A\<^sup>\) \ (B\<^sup>\)" proof show c0:"dim_row ((A \ B)\<^sup>\) = dim_row ((A\<^sup>\) \ (B\<^sup>\))" by simp show c1:"dim_col ((A \ B)\<^sup>\) = dim_col ((A\<^sup>\) \ (B\<^sup>\))" by simp show "\i j. i < dim_row ((A\<^sup>\) \ (B\<^sup>\)) \ j < dim_col ((A\<^sup>\) \ (B\<^sup>\)) \ ((A \ B)\<^sup>\) $$ (i, j) = ((A\<^sup>\) \ (B\<^sup>\)) $$ (i, j)" proof- fix i j assume a0:"i < dim_row ((A\<^sup>\) \ (B\<^sup>\))" and a1:"j < dim_col ((A\<^sup>\) \ (B\<^sup>\))" then have "(A \ B)\<^sup>\ $$ (i, j) = cnj((A \ B) $$ (j, i))" by (simp add: dagger_def) also have "\ = cnj(A $$ (j div dim_row(B), i div dim_col(B)) * B $$ (j mod dim_row(B), i mod dim_col(B)))" by (metis (mono_tags, lifting) a0 a1 c1 dim_row_tensor_mat dim_col_of_dagger dim_row_of_dagger index_tensor_mat less_nat_zero_code mult_not_zero neq0_conv) moreover have "((A\<^sup>\) \ (B\<^sup>\)) $$ (i, j) = (A\<^sup>\) $$ (i div dim_col(B), j div dim_row(B)) * (B\<^sup>\) $$ (i mod dim_col(B), j mod dim_row(B))" - by (smt a0 a1 c1 dim_row_tensor_mat dim_col_of_dagger dim_row_of_dagger index_tensor_mat + by (smt (verit) a0 a1 c1 dim_row_tensor_mat dim_col_of_dagger dim_row_of_dagger index_tensor_mat less_nat_zero_code mult_eq_0_iff neq0_conv) moreover have "(B\<^sup>\) $$ (i mod dim_col(B), j mod dim_row(B)) = cnj(B $$ (j mod dim_row(B), i mod dim_col(B)))" proof- have "i mod dim_col(B) < dim_col(B)" using a0 gr_implies_not_zero mod_div_trivial by fastforce moreover have "j mod dim_row(B) < dim_row(B)" using a1 gr_implies_not_zero mod_div_trivial by fastforce ultimately show ?thesis by (simp add: dagger_def) qed moreover have "(A\<^sup>\) $$ (i div dim_col(B), j div dim_row(B)) = cnj(A $$ (j div dim_row(B), i div dim_col(B)))" proof- have "i div dim_col(B) < dim_col(A)" using a0 dagger_def by (simp add: less_mult_imp_div_less) moreover have "j div dim_row(B) < dim_row(A)" using a1 dagger_def by (simp add: less_mult_imp_div_less) ultimately show ?thesis by (simp add: dagger_def) qed ultimately show "((A \ B)\<^sup>\) $$ (i, j) = ((A\<^sup>\) \ (B\<^sup>\)) $$ (i, j)" by simp qed qed locale quantum_machine = fixes n:: nat and s:: "complex Matrix.vec" and U:: "complex Matrix.mat" assumes dim_vec [simp]: "dim_vec s = 2^n" and dim_col [simp]: "dim_col U = 2^n * 2^n" and square [simp]: "square_mat U" and unitary [simp]: "unitary U" lemma inner_prod_of_unit_vec: fixes n i:: nat assumes "i < n" shows "\unit_vec n i| unit_vec n i\ = 1" by (auto simp add: inner_prod_def unit_vec_def) (simp add: assms sum.cong[of "{0..j. cnj (if j = i then 1 else 0) * (if j = i then 1 else 0)" "\j. (if j = i then 1 else 0)"]) theorem (in quantum_machine) no_cloning: assumes [simp]: "dim_vec v = 2^n" and [simp]: "dim_vec w = 2^n" and cloning1: "\s. U * ( |v\ \ |s\) = |v\ \ |v\" and cloning2: "\s. U * ( |w\ \ |s\) = |w\ \ |w\" and "\v|v\ = 1" and "\w|w\ = 1" shows "v = w \ \v|w\ = 0" proof- define s:: "complex Matrix.vec" where d0:"s = unit_vec (2^n) 0" have f0:"\|v\| \ \|s\| = (( |v\ \ |s\)\<^sup>\)" using hermite_cnj_of_tensor[of "|v\" "|s\"] bra_def dagger_def ket_vec_def by simp moreover have f1:"( |v\ \ |v\)\<^sup>\ * ( |w\ \ |w\) = (\|v\| \ \|s\| ) * ( |w\ \ |s\)" proof- have "(U * ( |v\ \ |s\))\<^sup>\ = (\|v\| \ \|s\| ) * (U\<^sup>\)" using dagger_of_prod[of "U" "|v\ \ |s\"] f0 d0 by (simp add: ket_vec_def) then have "(U * ( |v\ \ |s\))\<^sup>\ * U * ( |w\ \ |s\) = (\|v\| \ \|s\| ) * (U\<^sup>\) * U * ( |w\ \ |s\)" by simp moreover have "(U * ( |v\ \ |s\))\<^sup>\ * U * ( |w\ \ |s\) = (( |v\ \ |v\)\<^sup>\) * ( |w\ \ |w\)" - using assms(2-4) d0 unit_vec_def by (smt Matrix.dim_vec assoc_mult_mat carrier_mat_triv dim_row_mat(1) + using assms(2-4) d0 unit_vec_def by (smt (verit) Matrix.dim_vec assoc_mult_mat carrier_mat_triv dim_row_mat(1) dim_row_tensor_mat dim_col_of_dagger index_mult_mat(2) ket_vec_def square square_mat.elims(2)) moreover have "(U\<^sup>\) * U = 1\<^sub>m (2^n * 2^n)" using unitary_def dim_col unitary by simp moreover have "(\|v\| \ \|s\| ) * (U\<^sup>\) * U = (\|v\| \ \|s\| ) * ((U\<^sup>\) * U)" - using d0 assms(1) unit_vec_def by (smt Matrix.dim_vec assoc_mult_mat carrier_mat_triv dim_row_mat(1) + using d0 assms(1) unit_vec_def by (smt (verit) Matrix.dim_vec assoc_mult_mat carrier_mat_triv dim_row_mat(1) dim_row_tensor_mat f0 dim_col_of_dagger dim_row_of_dagger ket_vec_def local.dim_col) moreover have "(\|v\| \ \|s\| ) * 1\<^sub>m (2^n * 2^n) = (\|v\| \ \|s\| )" using f0 ket_vec_def d0 by simp ultimately show ?thesis by simp qed then have f2:"(\|v\| * |w\) \ (\|v\| * |w\) = (\|v\| * |w\) \ (\|s\| * |s\)" proof- have "\|v\| \ \|v\| = (( |v\ \ |v\)\<^sup>\)" using hermite_cnj_of_tensor[of "|v\" "|v\"] bra_def dagger_def ket_vec_def by simp then show ?thesis using f1 d0 by (simp add: bra_def mult_distr_tensor ket_vec_def) qed then have "\v|w\ * \v|w\ = \v|w\ * \s|s\" proof- have "((\|v\| * |w\) \ (\|v\| * |w\)) $$ (0,0) = \v|w\ * \v|w\" using assms inner_prod_with_times_mat[of "v" "w"] by (simp add: bra_def ket_vec_def) moreover have "((\|v\| * |w\) \ (\|s\| * |s\)) $$ (0,0) = \v|w\ * \s|s\" using inner_prod_with_times_mat[of "v" "w"] inner_prod_with_times_mat[of "s" "s"] by(simp add: bra_def ket_vec_def) ultimately show ?thesis using f2 by auto qed then have "\v|w\ = 0 \ \v|w\ = \s|s\" by (simp add: mult_left_cancel) moreover have "\s|s\ = 1" by(simp add: d0 inner_prod_of_unit_vec) ultimately show ?thesis using assms(1,2,5,6) by auto qed end \ No newline at end of file diff --git a/thys/Isabelle_Marries_Dirac/Quantum.thy b/thys/Isabelle_Marries_Dirac/Quantum.thy --- a/thys/Isabelle_Marries_Dirac/Quantum.thy +++ b/thys/Isabelle_Marries_Dirac/Quantum.thy @@ -1,1382 +1,1381 @@ (* Authors: Anthony Bordg, University of Cambridge, apdb3@cam.ac.uk Yijun He, University of Cambridge, yh403@cam.ac.uk with contributions by Hanna Lachnitt *) section \Qubits and Quantum Gates\ theory Quantum imports Jordan_Normal_Form.Matrix "HOL-Library.Nonpos_Ints" Basics Binary_Nat begin subsection \Qubits\ text\In this theory @{text cpx} stands for @{text complex}.\ definition cpx_vec_length :: "complex vec \ real" ("\_\") where "cpx_vec_length v \ sqrt(\i2)" lemma cpx_length_of_vec_of_list [simp]: "\vec_of_list l\ = sqrt(\i2)" by (auto simp: cpx_vec_length_def vec_of_list_def vec_of_list_index) (metis (no_types, lifting) dim_vec_of_list sum.cong vec_of_list.abs_eq vec_of_list_index) lemma norm_vec_index_unit_vec_is_0 [simp]: assumes "j < n" and "j \ i" shows "cmod ((unit_vec n i) $ j) = 0" using assms by (simp add: unit_vec_def) lemma norm_vec_index_unit_vec_is_1 [simp]: assumes "j < n" and "j = i" shows "cmod ((unit_vec n i) $ j) = 1" proof - have f:"(unit_vec n i) $ j = 1" using assms by simp thus ?thesis by (simp add: f cmod_def) qed lemma unit_cpx_vec_length [simp]: assumes "i < n" shows "\unit_vec n i\ = 1" proof - have "(\j2) = (\j = 1" using assms by simp finally have "sqrt (\j2) = 1" by simp thus ?thesis using cpx_vec_length_def by simp qed lemma smult_vec_length [simp]: assumes "x \ 0" shows "\complex_of_real(x) \\<^sub>v v\ = x * \v\" proof- have "(\i::nat.(cmod (complex_of_real x * v $ i))\<^sup>2) = (\i::nat. (cmod (v $ i))\<^sup>2 * x\<^sup>2)" by (auto simp: norm_mult power_mult_distrib) then have "(\i2) = (\i2 * x\<^sup>2)" by meson moreover have "(\i2 * x\<^sup>2) = x\<^sup>2 * (\i2)" by (metis (no_types) mult.commute sum_distrib_right) moreover have "sqrt(x\<^sup>2 * (\i2)) = sqrt(x\<^sup>2) * sqrt (\i2)" using real_sqrt_mult by blast ultimately show ?thesis by(simp add: cpx_vec_length_def assms) qed locale state = fixes n:: nat and v:: "complex mat" assumes is_column [simp]: "dim_col v = 1" and dim_row [simp]: "dim_row v = 2^n" and is_normal [simp]: "\col v 0\ = 1" text\ Below the natural number n codes for the dimension of the complex vector space whose elements of norm 1 we call states. \ lemma unit_vec_of_right_length_is_state [simp]: assumes "i < 2^n" shows "unit_vec (2^n) i \ {v| n v::complex vec. dim_vec v = 2^n \ \v\ = 1}" proof- have "dim_vec (unit_vec (2^n) i) = 2^n" by simp moreover have "\unit_vec (2^n) i\ = 1" using assms by simp ultimately show ?thesis by simp qed definition state_qbit :: "nat \ complex vec set" where "state_qbit n \ {v| v:: complex vec. dim_vec v = 2^n \ \v\ = 1}" lemma (in state) state_to_state_qbit [simp]: shows "col v 0 \ state_qbit n" using state_def state_qbit_def by simp subsection "The Hermitian Conjugation" text \The Hermitian conjugate of a complex matrix is the complex conjugate of its transpose. \ definition dagger :: "complex mat \ complex mat" ("_\<^sup>\") where "M\<^sup>\ \ mat (dim_col M) (dim_row M) (\(i,j). cnj(M $$ (j,i)))" text \We introduce the type of complex square matrices.\ typedef cpx_sqr_mat = "{M | M::complex mat. square_mat M}" proof- have "square_mat (1\<^sub>m n)" for n using one_mat_def by simp thus ?thesis by blast qed definition cpx_sqr_mat_to_cpx_mat :: "cpx_sqr_mat => complex mat" where "cpx_sqr_mat_to_cpx_mat M \ Rep_cpx_sqr_mat M" text \ We introduce a coercion from the type of complex square matrices to the type of complex matrices. \ declare [[coercion cpx_sqr_mat_to_cpx_mat]] lemma dim_row_of_dagger [simp]: "dim_row (M\<^sup>\) = dim_col M" using dagger_def by simp lemma dim_col_of_dagger [simp]: "dim_col (M\<^sup>\) = dim_row M" using dagger_def by simp lemma col_of_dagger [simp]: assumes "j < dim_row M" shows "col (M\<^sup>\) j = vec (dim_col M) (\i. cnj (M $$ (j,i)))" using assms col_def dagger_def by simp lemma row_of_dagger [simp]: assumes "i < dim_col M" shows "row (M\<^sup>\) i = vec (dim_row M) (\j. cnj (M $$ (j,i)))" using assms row_def dagger_def by simp lemma dagger_of_dagger_is_id: fixes M :: "complex Matrix.mat" shows "(M\<^sup>\)\<^sup>\ = M" proof show "dim_row ((M\<^sup>\)\<^sup>\) = dim_row M" by simp show "dim_col ((M\<^sup>\)\<^sup>\) = dim_col M" by simp fix i j assume a0:"i < dim_row M" and a1:"j < dim_col M" then show "(M\<^sup>\)\<^sup>\ $$ (i,j) = M $$ (i,j)" proof- show ?thesis using dagger_def a0 a1 by auto qed qed lemma dagger_of_sqr_is_sqr [simp]: "square_mat ((M::cpx_sqr_mat)\<^sup>\)" proof- have "square_mat M" using cpx_sqr_mat_to_cpx_mat_def Rep_cpx_sqr_mat by simp then have "dim_row M = dim_col M" by simp then have "dim_col (M\<^sup>\) = dim_row (M\<^sup>\)" by simp thus "square_mat (M\<^sup>\)" by simp qed lemma dagger_of_id_is_id [simp]: "(1\<^sub>m n)\<^sup>\ = 1\<^sub>m n" using dagger_def one_mat_def by auto subsection "Unitary Matrices and Quantum Gates" definition unitary :: "complex mat \ bool" where "unitary M \ (M\<^sup>\) * M = 1\<^sub>m (dim_col M) \ M * (M\<^sup>\) = 1\<^sub>m (dim_row M)" lemma id_is_unitary [simp]: "unitary (1\<^sub>m n)" by (simp add: unitary_def) locale gate = fixes n:: nat and A:: "complex mat" assumes dim_row [simp]: "dim_row A = 2^n" and square_mat [simp]: "square_mat A" and unitary [simp]: "unitary A" text \ We prove that a quantum gate is invertible and its inverse is given by its Hermitian conjugate. \ lemma mat_unitary_mat [intro]: assumes "unitary M" shows "inverts_mat M (M\<^sup>\)" using assms by (simp add: unitary_def inverts_mat_def) lemma unitary_mat_mat [intro]: assumes "unitary M" shows "inverts_mat (M\<^sup>\) M" using assms by (simp add: unitary_def inverts_mat_def) lemma (in gate) gate_is_inv: "invertible_mat A" using square_mat unitary invertible_mat_def by blast subsection "Relations Between Complex Conjugation, Hermitian Conjugation, Transposition and Unitarity" notation transpose_mat ("(_\<^sup>t)") lemma col_tranpose [simp]: assumes "dim_row M = n" and "i < n" shows "col (M\<^sup>t) i = row M i" proof show "dim_vec (col (M\<^sup>t) i) = dim_vec (row M i)" by (simp add: row_def col_def transpose_mat_def) next show "\j. j < dim_vec (row M i) \ col M\<^sup>t i $ j = row M i $ j" using assms by (simp add: transpose_mat_def) qed lemma row_transpose [simp]: assumes "dim_col M = n" and "i < n" shows "row (M\<^sup>t) i = col M i" using assms by simp definition cpx_mat_cnj :: "complex mat \ complex mat" ("(_\<^sup>\)") where "cpx_mat_cnj M \ mat (dim_row M) (dim_col M) (\(i,j). cnj (M $$ (i,j)))" lemma cpx_mat_cnj_id [simp]: "(1\<^sub>m n)\<^sup>\ = 1\<^sub>m n" by (auto simp: cpx_mat_cnj_def) lemma cpx_mat_cnj_cnj [simp]: "(M\<^sup>\)\<^sup>\ = M" by (auto simp: cpx_mat_cnj_def) lemma dim_row_of_cjn_prod [simp]: "dim_row ((M\<^sup>\) * (N\<^sup>\)) = dim_row M" by (simp add: cpx_mat_cnj_def) lemma dim_col_of_cjn_prod [simp]: "dim_col ((M\<^sup>\) * (N\<^sup>\)) = dim_col N" by (simp add: cpx_mat_cnj_def) lemma cpx_mat_cnj_prod: assumes "dim_col M = dim_row N" shows "(M * N)\<^sup>\ = (M\<^sup>\) * (N\<^sup>\)" proof show "dim_row (M * N)\<^sup>\ = dim_row ((M\<^sup>\) * (N\<^sup>\))" by (simp add: cpx_mat_cnj_def) next show "dim_col ((M * N)\<^sup>\) = dim_col ((M\<^sup>\) * (N\<^sup>\))" by (simp add: cpx_mat_cnj_def) next fix i j::nat assume a1:"i < dim_row ((M\<^sup>\) * (N\<^sup>\))" and a2:"j < dim_col ((M\<^sup>\) * (N\<^sup>\))" then have "(M * N)\<^sup>\ $$ (i,j) = cnj (\k<(dim_row N). M $$ (i,k) * N $$ (k,j))" using assms cpx_mat_cnj_def index_mat times_mat_def scalar_prod_def row_def col_def dim_row_of_cjn_prod dim_col_of_cjn_prod - by (smt case_prod_conv dim_col index_mult_mat(2) index_mult_mat(3) index_vec lessThan_atLeast0 - lessThan_iff sum.cong) + by simp also have "\ = (\k<(dim_row N). cnj(M $$ (i,k)) * cnj(N $$ (k,j)))" by simp also have "((M\<^sup>\) * (N\<^sup>\)) $$ (i,j) = (\k<(dim_row N). cnj(M $$ (i,k)) * cnj(N $$ (k,j)))" using assms a1 a2 cpx_mat_cnj_def index_mat times_mat_def scalar_prod_def row_def col_def by (smt case_prod_conv dim_col dim_col_mat(1) dim_row_mat(1) index_vec lessThan_atLeast0 lessThan_iff sum.cong) finally show "(M * N)\<^sup>\ $$ (i, j) = ((M\<^sup>\) * (N\<^sup>\)) $$ (i, j)" by simp qed lemma transpose_of_prod: fixes M N::"complex Matrix.mat" assumes "dim_col M = dim_row N" shows "(M * N)\<^sup>t = N\<^sup>t * (M\<^sup>t)" proof fix i j::nat assume a0: "i < dim_row (N\<^sup>t * (M\<^sup>t))" and a1: "j < dim_col (N\<^sup>t * (M\<^sup>t))" then have "(M * N)\<^sup>t $$ (i,j) = (M * N) $$ (j,i)" by auto also have "... = (\kt. M $$ (j,k) * N $$ (k,i))" using assms a0 a1 by auto also have "... = (\kt. N $$ (k,i) * M $$ (j,k))" by (simp add: semiring_normalization_rules(7)) also have "... = (\kt. ((N\<^sup>t) $$ (i,k)) * (M\<^sup>t) $$ (k,j))" using assms a0 a1 by auto finally show "((M * N)\<^sup>t) $$ (i,j) = (N\<^sup>t * (M\<^sup>t)) $$ (i,j)" using assms a0 a1 by auto next show "dim_row ((M * N)\<^sup>t) = dim_row (N\<^sup>t * (M\<^sup>t))" by auto next show "dim_col ((M * N)\<^sup>t) = dim_col (N\<^sup>t * (M\<^sup>t))" by auto qed lemma transpose_cnj_is_dagger [simp]: "(M\<^sup>t)\<^sup>\ = (M\<^sup>\)" proof show f1:"dim_row ((M\<^sup>t)\<^sup>\) = dim_row (M\<^sup>\)" by (simp add: cpx_mat_cnj_def transpose_mat_def dagger_def) next show f2:"dim_col ((M\<^sup>t)\<^sup>\) = dim_col (M\<^sup>\)" by (simp add: cpx_mat_cnj_def transpose_mat_def dagger_def) next fix i j::nat assume "i < dim_row M\<^sup>\" and "j < dim_col M\<^sup>\" then show "M\<^sup>t\<^sup>\ $$ (i, j) = M\<^sup>\ $$ (i, j)" by (simp add: cpx_mat_cnj_def transpose_mat_def dagger_def) qed lemma cnj_transpose_is_dagger [simp]: "(M\<^sup>\)\<^sup>t = (M\<^sup>\)" proof show "dim_row ((M\<^sup>\)\<^sup>t) = dim_row (M\<^sup>\)" by (simp add: transpose_mat_def cpx_mat_cnj_def dagger_def) next show "dim_col ((M\<^sup>\)\<^sup>t) = dim_col (M\<^sup>\)" by (simp add: transpose_mat_def cpx_mat_cnj_def dagger_def) next fix i j::nat assume "i < dim_row M\<^sup>\" and "j < dim_col M\<^sup>\" then show "M\<^sup>\\<^sup>t $$ (i, j) = M\<^sup>\ $$ (i, j)" by (simp add: transpose_mat_def cpx_mat_cnj_def dagger_def) qed lemma dagger_of_transpose_is_cnj [simp]: "(M\<^sup>t)\<^sup>\ = (M\<^sup>\)" by (metis transpose_transpose transpose_cnj_is_dagger) lemma dagger_of_prod: fixes M N::"complex Matrix.mat" assumes "dim_col M = dim_row N" shows "(M * N)\<^sup>\ = N\<^sup>\ * (M\<^sup>\)" proof- have "(M * N)\<^sup>\ = ((M * N)\<^sup>\)\<^sup>t" by auto also have "... = ((M\<^sup>\) * (N\<^sup>\))\<^sup>t" using assms cpx_mat_cnj_prod by auto also have "... = (N\<^sup>\)\<^sup>t * ((M\<^sup>\)\<^sup>t)" using assms transpose_of_prod by (metis cnj_transpose_is_dagger dim_col_of_dagger dim_row_of_dagger index_transpose_mat(2) index_transpose_mat(3)) finally show "(M * N)\<^sup>\ = N\<^sup>\ * (M\<^sup>\)" by auto qed text \The product of two quantum gates is a quantum gate.\ lemma prod_of_gate_is_gate: assumes "gate n G1" and "gate n G2" shows "gate n (G1 * G2)" proof show "dim_row (G1 * G2) = 2^n" using assms by (simp add: gate_def) next show "square_mat (G1 * G2)" using assms gate.dim_row gate.square_mat by simp next show "unitary (G1 * G2)" proof- have "((G1 * G2)\<^sup>\) * (G1 * G2) = 1\<^sub>m (dim_col (G1 * G2))" proof- have f0: "G1 \ carrier_mat (2^n) (2^n) \ G2 \ carrier_mat (2^n) (2^n) \ G1\<^sup>\ \ carrier_mat (2^n) (2^n) \ G2\<^sup>\ \ carrier_mat (2^n) (2^n) \ G1 * G2 \ carrier_mat (2^n) (2^n)" using assms gate.dim_row gate.square_mat by auto have "((G1 * G2)\<^sup>\) * (G1 * G2) = ((G2\<^sup>\) * (G1\<^sup>\)) * (G1 * G2)" using assms dagger_of_prod gate.dim_row gate.square_mat by simp also have "... = (G2\<^sup>\) * ((G1\<^sup>\) * (G1 * G2))" using assms f0 by auto also have "... = (G2\<^sup>\) * (((G1\<^sup>\) * G1) * G2)" using assms f0 f0 by auto also have "... = (G2\<^sup>\) * ((1\<^sub>m (dim_col G1)) * G2)" using gate.unitary[of n G1] assms unitary_def[of G1] by simp also have "... = (G2\<^sup>\) * ((1\<^sub>m (dim_col G2)) * G2)" using assms f0 by (metis carrier_matD(2)) also have "... = (G2\<^sup>\) * G2" using f0 by (metis carrier_matD(2) left_mult_one_mat) finally show "((G1 * G2)\<^sup>\) * (G1 * G2) = 1\<^sub>m (dim_col (G1 * G2))" using assms gate.unitary unitary_def by simp qed moreover have "(G1 * G2) * ((G1 * G2)\<^sup>\) = 1\<^sub>m (dim_row (G1 * G2))" using assms calculation - by (smt carrier_matI dim_col_of_dagger dim_row_of_dagger gate.dim_row gate.square_mat index_mult_mat(2) index_mult_mat(3) + by (smt (verit) carrier_matI dim_col_of_dagger dim_row_of_dagger gate.dim_row gate.square_mat index_mult_mat(2) index_mult_mat(3) mat_mult_left_right_inverse square_mat.elims(2)) ultimately show ?thesis using unitary_def by simp qed qed lemma left_inv_of_unitary_transpose [simp]: assumes "unitary U" shows "(U\<^sup>t)\<^sup>\ * (U\<^sup>t) = 1\<^sub>m(dim_row U)" proof - have "dim_col U = dim_row ((U\<^sup>t)\<^sup>\)" by simp then have "(U * ((U\<^sup>t)\<^sup>\))\<^sup>\ = (U\<^sup>\) * (U\<^sup>t)" using cpx_mat_cnj_prod cpx_mat_cnj_cnj by presburger also have "\ = (U\<^sup>t)\<^sup>\ * (U\<^sup>t)" by simp finally show ?thesis using assms by (metis transpose_cnj_is_dagger cpx_mat_cnj_id unitary_def) qed lemma right_inv_of_unitary_transpose [simp]: assumes "unitary U" shows "U\<^sup>t * ((U\<^sup>t)\<^sup>\) = 1\<^sub>m(dim_col U)" proof - have "dim_col ((U\<^sup>t)\<^sup>\) = dim_row U" by simp then have "U\<^sup>t * ((U\<^sup>t)\<^sup>\) = (((U\<^sup>t)\<^sup>\ * U)\<^sup>\)" using cpx_mat_cnj_cnj cpx_mat_cnj_prod dagger_of_transpose_is_cnj by presburger also have "\ = (U\<^sup>\ * U)\<^sup>\" by simp finally show ?thesis using assms by (metis cpx_mat_cnj_id unitary_def) qed lemma transpose_of_unitary_is_unitary [simp]: assumes "unitary U" shows "unitary (U\<^sup>t)" using unitary_def assms left_inv_of_unitary_transpose right_inv_of_unitary_transpose by simp subsection "The Inner Product" text \We introduce a coercion between complex vectors and (column) complex matrices.\ definition ket_vec :: "complex vec \ complex mat" ("|_\") where "|v\ \ mat (dim_vec v) 1 (\(i,j). v $ i)" lemma ket_vec_index [simp]: assumes "i < dim_vec v" shows "|v\ $$ (i,0) = v $ i" using assms ket_vec_def by simp lemma ket_vec_col [simp]: "col |v\ 0 = v" by (auto simp: col_def ket_vec_def) lemma smult_ket_vec [simp]: "|x \\<^sub>v v\ = x \\<^sub>m |v\" by (auto simp: ket_vec_def) lemma smult_vec_length_bis [simp]: assumes "x \ 0" shows "\col (complex_of_real(x) \\<^sub>m |v\) 0\ = x * \v\" using assms smult_ket_vec smult_vec_length ket_vec_col by metis declare [[coercion ket_vec]] definition row_vec :: "complex vec \ complex mat" where "row_vec v \ mat 1 (dim_vec v) (\(i,j). v $ j)" definition bra_vec :: "complex vec \ complex mat" where "bra_vec v \ (row_vec v)\<^sup>\" lemma row_bra_vec [simp]: "row (bra_vec v) 0 = vec (dim_vec v) (\i. cnj(v $ i))" by (auto simp: row_def bra_vec_def cpx_mat_cnj_def row_vec_def) text \We introduce a definition called @{term "bra"} to see a vector as a column matrix.\ definition bra :: "complex mat \ complex mat" ("\_|") where "\v| \ mat 1 (dim_row v) (\(i,j). cnj(v $$ (j,i)))" text \The relation between @{term "bra"}, @{term "bra_vec"} and @{term "ket_vec"} is given as follows.\ lemma bra_bra_vec [simp]: "bra (ket_vec v) = bra_vec v" by (auto simp: bra_def ket_vec_def bra_vec_def cpx_mat_cnj_def row_vec_def) lemma row_bra [simp]: fixes v::"complex vec" shows "row \v| 0 = vec (dim_vec v) (\i. cnj (v $ i))" by simp text \We introduce the inner product of two complex vectors in @{text "\\<^sup>n"}.\ definition inner_prod :: "complex vec \ complex vec \ complex" ("\_|_\") where "inner_prod u v \ \ i \ {0..< dim_vec v}. cnj(u $ i) * (v $ i)" lemma inner_prod_with_row_bra_vec [simp]: assumes "dim_vec u = dim_vec v" shows "\u|v\ = row (bra_vec u) 0 \ v" using assms inner_prod_def scalar_prod_def row_bra_vec index_vec - by (smt lessThan_atLeast0 lessThan_iff sum.cong) + by (smt (verit) lessThan_atLeast0 lessThan_iff sum.cong) lemma inner_prod_with_row_bra_vec_col_ket_vec [simp]: assumes "dim_vec u = dim_vec v" shows "\u|v\ = (row \u| 0) \ (col |v\ 0)" using assms by (simp add: inner_prod_def scalar_prod_def) lemma inner_prod_with_times_mat [simp]: assumes "dim_vec u = dim_vec v" shows "\u|v\ = (\u| * |v\) $$ (0,0)" using assms inner_prod_with_row_bra_vec_col_ket_vec by (simp add: inner_prod_def times_mat_def ket_vec_def bra_def) lemma orthogonal_unit_vec [simp]: assumes "i < n" and "j < n" and "i \ j" shows "\unit_vec n i|unit_vec n j\ = 0" proof- have "\unit_vec n i|unit_vec n j\ = unit_vec n i \ unit_vec n j" using assms unit_vec_def inner_prod_def scalar_prod_def - by (smt complex_cnj_zero index_unit_vec(3) index_vec inner_prod_with_row_bra_vec row_bra_vec + by (smt (verit) complex_cnj_zero index_unit_vec(3) index_vec inner_prod_with_row_bra_vec row_bra_vec scalar_prod_right_unit) thus ?thesis using assms scalar_prod_def unit_vec_def by simp qed text \We prove that our inner product is linear in its second argument.\ lemma vec_index_is_linear [simp]: assumes "dim_vec u = dim_vec v" and "j < dim_vec u" shows "(k \\<^sub>v u + l \\<^sub>v v) $ j = k * (u $ j) + l * (v $ j)" using assms vec_index_def smult_vec_def plus_vec_def by simp lemma inner_prod_is_linear [simp]: fixes u::"complex vec" and v::"nat \ complex vec" and l::"nat \ complex" assumes "\i\{0, 1}. dim_vec u = dim_vec (v i)" shows "\u|l 0 \\<^sub>v v 0 + l 1 \\<^sub>v v 1\ = (\i\1. l i * \u|v i\)" proof - have f1:"dim_vec (l 0 \\<^sub>v v 0 + l 1 \\<^sub>v v 1) = dim_vec u" using assms by simp then have "\u|l 0 \\<^sub>v v 0 + l 1 \\<^sub>v v 1\ = (\i\{0 ..< dim_vec u}. cnj (u $ i) * ((l 0 \\<^sub>v v 0 + l 1 \\<^sub>v v 1) $ i))" by (simp add: inner_prod_def) also have "\ = (\i\{0 ..< dim_vec u}. cnj (u $ i) * (l 0 * v 0 $ i + l 1 * v 1 $ i))" using assms by simp also have "\ = l 0 * (\i\{0 ..< dim_vec u}. cnj(u $ i) * (v 0 $ i)) + l 1 * (\i\{0 ..< dim_vec u}. cnj(u $ i) * (v 1 $ i))" by (auto simp: algebra_simps) (simp add: sum.distrib sum_distrib_left) also have "\ = l 0 * \u|v 0\ + l 1 * \u|v 1\" using assms inner_prod_def by auto finally show ?thesis by simp qed lemma inner_prod_cnj: assumes "dim_vec u = dim_vec v" shows "\v|u\ = cnj (\u|v\)" by (simp add: assms inner_prod_def algebra_simps) lemma inner_prod_with_itself_Im [simp]: "Im (\u|u\) = 0" using inner_prod_cnj by (metis Reals_cnj_iff complex_is_Real_iff) lemma inner_prod_with_itself_real [simp]: "\u|u\ \ \" using inner_prod_with_itself_Im by (simp add: complex_is_Real_iff) lemma inner_prod_with_itself_eq0 [simp]: assumes "u = 0\<^sub>v (dim_vec u)" shows "\u|u\ = 0" using assms inner_prod_def zero_vec_def - by (smt atLeastLessThan_iff complex_cnj_zero index_zero_vec(1) mult_zero_left sum.neutral) + by (smt (verit) atLeastLessThan_iff complex_cnj_zero index_zero_vec(1) mult_zero_left sum.neutral) lemma inner_prod_with_itself_Re: "Re (\u|u\) \ 0" proof - have "Re (\u|u\) = (\i = (\i2 + (Im (u $ i))\<^sup>2)" using complex_mult_cnj by (metis (no_types, lifting) Re_complex_of_real semiring_normalization_rules(7)) ultimately show "Re (\u|u\) \ 0" by (simp add: sum_nonneg) qed lemma inner_prod_with_itself_nonneg_reals: fixes u::"complex vec" shows "\u|u\ \ nonneg_Reals" using inner_prod_with_itself_real inner_prod_with_itself_Re complex_nonneg_Reals_iff inner_prod_with_itself_Im by auto lemma inner_prod_with_itself_Re_non0: assumes "u \ 0\<^sub>v (dim_vec u)" shows "Re (\u|u\) > 0" proof - obtain i where a1:"i < dim_vec u" and "u $ i \ 0" using assms zero_vec_def by (metis dim_vec eq_vecI index_zero_vec(1)) then have f1:"Re (cnj (u $ i) * (u $ i)) > 0" by (metis Re_complex_of_real complex_mult_cnj complex_neq_0 mult.commute) moreover have f2:"Re (\u|u\) = (\ii 0" using complex_mult_cnj by simp ultimately show ?thesis using a1 inner_prod_def lessThan_iff by (metis (no_types, lifting) finite_lessThan sum_pos2) qed lemma inner_prod_with_itself_nonneg_reals_non0: assumes "u \ 0\<^sub>v (dim_vec u)" shows "\u|u\ \ 0" using assms inner_prod_with_itself_Re_non0 by fastforce lemma cpx_vec_length_inner_prod [simp]: "\v\\<^sup>2 = \v|v\" proof - have "\v\\<^sup>2 = (\i2)" using cpx_vec_length_def complex_of_real_def by (metis (no_types, lifting) real_sqrt_power real_sqrt_unique sum_nonneg zero_le_power2) also have "\ = (\iv|v\ = \v\" using inner_prod_with_itself_Re inner_prod_with_itself_Im csqrt_of_real_nonneg cpx_vec_length_def by (metis (no_types, lifting) Re_complex_of_real cpx_vec_length_inner_prod real_sqrt_ge_0_iff real_sqrt_unique sum_nonneg zero_le_power2) subsection "Unitary Matrices and Length-Preservation" subsubsection "Unitary Matrices are Length-Preserving" text \The bra-vector @{text "\A * v|"} is given by @{text "\v| * A\<^sup>\"}\ lemma dagger_of_ket_is_bra: fixes v:: "complex vec" shows "( |v\ )\<^sup>\ = \v|" by (simp add: bra_def dagger_def ket_vec_def) lemma bra_mat_on_vec: fixes v::"complex vec" and A::"complex mat" assumes "dim_col A = dim_vec v" shows "\A * v| = \v| * (A\<^sup>\)" proof show "dim_row \A * v| = dim_row (\v| * (A\<^sup>\))" by (simp add: bra_def times_mat_def) next show "dim_col \A * v| = dim_col (\v| * (A\<^sup>\))" by (simp add: bra_def times_mat_def) next fix i j::nat assume a1:"i < dim_row (\v| * (A\<^sup>\))" and a2:"j < dim_col (\v| * (A\<^sup>\))" then have "cnj((A * v) $$ (j,0)) = cnj (row A j \ v)" using bra_def times_mat_def ket_vec_col ket_vec_def by simp also have f7:"\= (\i\{0 ..< dim_vec v}. cnj(v $ i) * cnj(A $$ (j,i)))" using row_def scalar_prod_def cnj_sum complex_cnj_mult mult.commute by (smt assms index_vec lessThan_atLeast0 lessThan_iff sum.cong) moreover have f8:"(row \v| 0) \ (col (A\<^sup>\) j) = vec (dim_vec v) (\i. cnj (v $ i)) \ vec (dim_col A) (\i. cnj (A $$ (j,i)))" using a2 by simp ultimately have "cnj((A * v) $$ (j,0)) = (row \v| 0) \ (col (A\<^sup>\) j)" using assms scalar_prod_def - by (smt dim_vec index_vec lessThan_atLeast0 lessThan_iff sum.cong) + by (smt (verit) dim_vec index_vec lessThan_atLeast0 lessThan_iff sum.cong) then have "\A * v| $$ (0,j) = (\v| * (A\<^sup>\)) $$ (0,j)" using bra_def times_mat_def a2 by simp thus "\A * |v\| $$ (i, j) = (\v| * (A\<^sup>\)) $$ (i, j)" using a1 by (simp add: times_mat_def bra_def) qed lemma mat_on_ket: fixes v:: "complex vec" and A:: "complex mat" assumes "dim_col A = dim_vec v" shows "A * |v\ = |col (A * v) 0\" using assms ket_vec_def by auto lemma dagger_of_mat_on_ket: fixes v:: "complex vec" and A :: "complex mat" assumes "dim_col A = dim_vec v" shows "(A * |v\ )\<^sup>\ = \v| * (A\<^sup>\)" using assms by (metis bra_mat_on_vec dagger_of_ket_is_bra mat_on_ket) definition col_fst :: "'a mat \ 'a vec" where "col_fst A = vec (dim_row A) (\ i. A $$ (i,0))" lemma col_fst_is_col [simp]: "col_fst M = col M 0" by (simp add: col_def col_fst_def) text \ We need to declare @{term "col_fst"} as a coercion from matrices to vectors in order to see a column matrix as a vector. \ declare [[coercion_delete ket_vec]] [[coercion col_fst]] lemma unit_vec_to_col: assumes "dim_col A = n" and "i < n" shows "col A i = A * |unit_vec n i\" proof show "dim_vec (col A i) = dim_vec (A * |unit_vec n i\)" using col_def times_mat_def by simp next fix j::nat assume "j < dim_vec (col_fst (A * |unit_vec n i\))" then show "col A i $ j = (A * |unit_vec n i\) $ j" using assms times_mat_def ket_vec_def - by (smt col_fst_is_col dim_col dim_col_mat(1) index_col index_mult_mat(1) index_mult_mat(2) + by (smt (verit) col_fst_is_col dim_col dim_col_mat(1) index_col index_mult_mat(1) index_mult_mat(2) index_row(1) ket_vec_col less_numeral_extra(1) scalar_prod_right_unit) qed lemma mult_ket_vec_is_ket_vec_of_mult: fixes A::"complex mat" and v::"complex vec" assumes "dim_col A = dim_vec v" shows "|A * |v\ \ = A * |v\" using assms ket_vec_def by (metis One_nat_def col_fst_is_col dim_col dim_col_mat(1) index_mult_mat(3) ket_vec_col less_Suc0 mat_col_eqI) lemma unitary_is_sq_length_preserving [simp]: assumes "unitary U" and "dim_vec v = dim_col U" shows "\U * |v\\\<^sup>2 = \v\\<^sup>2" proof - have "\U * |v\|U * |v\ \ = (\|v\| * (U\<^sup>\) * |U * |v\\) $$ (0,0)" using assms(2) bra_mat_on_vec by (metis inner_prod_with_times_mat mult_ket_vec_is_ket_vec_of_mult) then have "\U * |v\|U * |v\ \ = (\|v\| * (U\<^sup>\) * (U * |v\)) $$ (0,0)" using assms(2) mult_ket_vec_is_ket_vec_of_mult by simp moreover have f1:"dim_col \|v\| = dim_vec v" using ket_vec_def bra_def by simp moreover have "dim_row (U\<^sup>\) = dim_vec v" using assms(2) by simp ultimately have "\U * |v\|U * |v\ \ = (\|v\| * ((U\<^sup>\) * U) * |v\) $$ (0,0)" using assoc_mult_mat - by(smt carrier_mat_triv dim_row_mat(1) dagger_def ket_vec_def mat_carrier times_mat_def) + by(smt (verit, ccfv_threshold) carrier_mat_triv dim_row_mat(1) dagger_def ket_vec_def mat_carrier times_mat_def) then have "\U * |v\|U * |v\ \ = (\|v\| * |v\) $$ (0,0)" using assms f1 unitary_def by simp thus ?thesis using cpx_vec_length_inner_prod by(metis Re_complex_of_real inner_prod_with_times_mat) qed lemma col_ket_vec [simp]: assumes "dim_col M = 1" shows "|col M 0\ = M" using eq_matI assms ket_vec_def by auto lemma state_col_ket_vec: assumes "state 1 v" shows "state 1 |col v 0\" using assms by (simp add: state_def) lemma col_ket_vec_index [simp]: assumes "i < dim_row v" shows "|col v 0\ $$ (i,0) = v $$ (i,0)" using assms ket_vec_def by (simp add: col_def) lemma col_index_of_mat_col [simp]: assumes "dim_col v = 1" and "i < dim_row v" shows "col v 0 $ i = v $$ (i,0)" using assms by simp lemma unitary_is_sq_length_preserving_bis [simp]: assumes "unitary U" and "dim_row v = dim_col U" and "dim_col v = 1" shows "\col (U * v) 0\\<^sup>2 = \col v 0\\<^sup>2" proof - have "dim_vec (col v 0) = dim_col U" using assms(2) by simp then have "\col_fst (U * |col v 0\)\\<^sup>2 = \col v 0\\<^sup>2" using unitary_is_sq_length_preserving[of "U" "col v 0"] assms(1) by simp thus ?thesis using assms(3) by simp qed text \ A unitary matrix is length-preserving, i.e. it acts on a vector to produce another vector of the same length. \ lemma unitary_is_length_preserving_bis [simp]: fixes U::"complex mat" and v::"complex mat" assumes "unitary U" and "dim_row v = dim_col U" and "dim_col v = 1" shows "\col (U * v) 0\ = \col v 0\" using assms unitary_is_sq_length_preserving_bis by (metis cpx_vec_length_inner_prod inner_prod_csqrt of_real_hom.injectivity) lemma unitary_is_length_preserving [simp]: fixes U:: "complex mat" and v:: "complex vec" assumes "unitary U" and "dim_vec v = dim_col U" shows "\U * |v\\ = \v\" using assms unitary_is_sq_length_preserving by (metis cpx_vec_length_inner_prod inner_prod_csqrt of_real_hom.injectivity) subsubsection "Length-Preserving Matrices are Unitary" lemma inverts_mat_sym: fixes A B:: "complex mat" assumes "inverts_mat A B" and "dim_row B = dim_col A" and "square_mat B" shows "inverts_mat B A" proof- define n where d0:"n = dim_row B" have "A * B = 1\<^sub>m (dim_row A)" using assms(1) inverts_mat_def by auto moreover have "dim_col B = dim_col (A * B)" using times_mat_def by simp ultimately have "dim_col B = dim_row A" by simp then have c0:"A \ carrier_mat n n" using assms(2,3) d0 by auto have c1:"B \ carrier_mat n n" using assms(3) d0 by auto have f0:"A * B = 1\<^sub>m n" using inverts_mat_def c0 c1 assms(1) by auto have f1:"det B \ 0" proof assume "det B = 0" then have "\v. v \ carrier_vec n \ v \ 0\<^sub>v n \ B *\<^sub>v v = 0\<^sub>v n" using det_0_iff_vec_prod_zero assms(3) c1 by blast then obtain v where d1:"v \ carrier_vec n \ v \ 0\<^sub>v n \ B *\<^sub>v v = 0\<^sub>v n" by auto then have d2:"dim_vec v = n" by simp have "B * |v\ = |0\<^sub>v n\" proof show "dim_row (B * |v\) = dim_row |0\<^sub>v n\" using ket_vec_def d0 by simp next show "dim_col (B * |v\) = dim_col |0\<^sub>v n\" using ket_vec_def d0 by simp next fix i j assume "i < dim_row |0\<^sub>v n\" and "j < dim_col |0\<^sub>v n\" then have f2:"i < n \ j = 0" using ket_vec_def by simp moreover have "vec (dim_row B) (($) v) = v" using d0 d1 by auto moreover have "(B *\<^sub>v v) $ i = (\ia = 0..) $$ (i, j) = |0\<^sub>v n\ $$ (i, j)" using ket_vec_def d0 d1 times_mat_def mult_mat_vec_def by (auto simp add: scalar_prod_def) qed moreover have "|v\ \ carrier_mat n 1" using d2 ket_vec_def by simp ultimately have "(A * B) * |v\ = A * |0\<^sub>v n\" using c0 c1 by simp then have f3:"|v\ = A * |0\<^sub>v n\" using d2 f0 ket_vec_def by auto have "v = 0\<^sub>v n" proof show "dim_vec v = dim_vec (0\<^sub>v n)" using d2 by simp next fix i assume f4:"i < dim_vec (0\<^sub>v n)" then have "|v\ $$ (i,0) = v $ i" using d2 ket_vec_def by simp moreover have "(A * |0\<^sub>v n\) $$ (i, 0) = 0" using ket_vec_def times_mat_def scalar_prod_def f4 c0 by auto ultimately show "v $ i = 0\<^sub>v n $ i" using f3 f4 by simp qed then show False using d1 by simp qed have f5:"adj_mat B \ carrier_mat n n \ B * adj_mat B = det B \\<^sub>m 1\<^sub>m n" using c1 adj_mat by auto then have c2:"((1/det B) \\<^sub>m adj_mat B) \ carrier_mat n n" by simp have f6:"B * ((1/det B) \\<^sub>m adj_mat B) = 1\<^sub>m n" using c1 f1 f5 mult_smult_distrib[of "B"] by auto then have "A = (A * B) * ((1/det B) \\<^sub>m adj_mat B)" using c0 c1 c2 by simp then have "A = (1/det B) \\<^sub>m adj_mat B" using f0 c2 by auto then show ?thesis using c0 c1 f6 inverts_mat_def by auto qed lemma sum_of_unit_vec_length: fixes i j n:: nat and c:: complex assumes "i < n" and "j < n" and "i \ j" shows "\unit_vec n i + c \\<^sub>v unit_vec n j\\<^sup>2 = 1 + cnj(c) * c" proof- define v where d0:"v = unit_vec n i + c \\<^sub>v unit_vec n j" have "\kkv\\<^sup>2 = (\k = 0..\<^sub>v col A j = A * |unit_vec n i + c \\<^sub>v unit_vec n j\" proof show "dim_vec (col A i + c \\<^sub>v col A j) = dim_vec (col_fst (A * |unit_vec n i + c \\<^sub>v unit_vec n j\))" using assms(1) by auto next fix k assume "k < dim_vec (col_fst (A * |unit_vec n i + c \\<^sub>v unit_vec n j\))" then have f0:"k < dim_row A" using assms(1) by auto have "(col A i + c \\<^sub>v col A j) $ k = A $$ (k, i) + c * A $$ (k, j)" using f0 assms(1-3) by auto moreover have "(\xxxxx\<^sub>v col A j) $ k = col_fst (A * |unit_vec n i + c \\<^sub>v unit_vec n j\) $ k" using f0 assms(1-3) times_mat_def scalar_prod_def ket_vec_def by auto qed lemma inner_prod_is_sesquilinear: fixes u1 u2 v1 v2:: "complex vec" and c1 c2 c3 c4:: complex and n:: nat assumes "dim_vec u1 = n" and "dim_vec u2 = n" and "dim_vec v1 = n" and "dim_vec v2 = n" shows "\c1 \\<^sub>v u1 + c2 \\<^sub>v u2|c3 \\<^sub>v v1 + c4 \\<^sub>v v2\ = cnj (c1) * c3 * \u1|v1\ + cnj (c2) * c3 * \u2|v1\ + cnj (c1) * c4 * \u1|v2\ + cnj (c2) * c4 * \u2|v2\" proof- have "\c1 \\<^sub>v u1 + c2 \\<^sub>v u2|c3 \\<^sub>v v1 + c4 \\<^sub>v v2\ = c3 * \c1 \\<^sub>v u1 + c2 \\<^sub>v u2|v1\ + c4 * \c1 \\<^sub>v u1 + c2 \\<^sub>v u2|v2\" using inner_prod_is_linear[of "c1 \\<^sub>v u1 + c2 \\<^sub>v u2" "\i. if i = 0 then v1 else v2" "\i. if i = 0 then c3 else c4"] assms by simp also have "... = c3 * cnj(\v1|c1 \\<^sub>v u1 + c2 \\<^sub>v u2\) + c4 * cnj(\v2|c1 \\<^sub>v u1 + c2 \\<^sub>v u2\)" using assms inner_prod_cnj[of "v1" "c1 \\<^sub>v u1 + c2 \\<^sub>v u2"] inner_prod_cnj[of "v2" "c1 \\<^sub>v u1 + c2 \\<^sub>v u2"] by simp also have "... = c3 * cnj(c1 * \v1|u1\ + c2 * \v1|u2\) + c4 * cnj(c1 * \v2|u1\ + c2 * \v2|u2\)" using inner_prod_is_linear[of "v1" "\i. if i = 0 then u1 else u2" "\i. if i = 0 then c1 else c2"] inner_prod_is_linear[of "v2" "\i. if i = 0 then u1 else u2" "\i. if i = 0 then c1 else c2"] assms by simp also have "... = c3 * (cnj(c1) * \u1|v1\ + cnj(c2) * \u2|v1\) + c4 * (cnj(c1) * \u1|v2\ + cnj(c2) * \u2|v2\)" using inner_prod_cnj[of "v1" "u1"] inner_prod_cnj[of "v1" "u2"] inner_prod_cnj[of "v2" "u1"] inner_prod_cnj[of "v2" "u2"] assms by simp finally show ?thesis by (auto simp add: algebra_simps) qed text \ A length-preserving matrix is unitary. So, unitary matrices are exactly the length-preserving matrices. \ lemma length_preserving_is_unitary: fixes U:: "complex mat" assumes "square_mat U" and "\v::complex vec. dim_vec v = dim_col U \ \U * |v\\ = \v\" shows "unitary U" proof- define n where "n = dim_col U" then have c0:"U \ carrier_mat n n" using assms(1) by auto then have c1:"U\<^sup>\ \ carrier_mat n n" using assms(1) dagger_def by auto have f0:"(U\<^sup>\) * U = 1\<^sub>m (dim_col U)" proof show "dim_row (U\<^sup>\ * U) = dim_row (1\<^sub>m (dim_col U))" using c0 by simp next show "dim_col (U\<^sup>\ * U) = dim_col (1\<^sub>m (dim_col U))" using c0 by simp next fix i j assume "i < dim_row (1\<^sub>m (dim_col U))" and "j < dim_col (1\<^sub>m (dim_col U))" then have a0:"i < n \ j < n" using c0 by simp have f1:"\l. l (\kcol U l\\<^sup>2 = (\kcol U l\\<^sup>2 = \v\\<^sup>2" using c0 d1 a1 assms(2) unit_vec_to_col by simp moreover have "\v\\<^sup>2 = 1" using d1 a1 cpx_vec_length_inner_prod by simp ultimately show "(\k j \ (\k j" define v1::"complex vec" where d1:"v1 = unit_vec n i + 1 \\<^sub>v unit_vec n j" define v2::"complex vec" where d2:"v2 = unit_vec n i + \ \\<^sub>v unit_vec n j" have "\v1\\<^sup>2 = 1 + cnj 1 * 1" using d1 a0 a2 sum_of_unit_vec_length by blast then have "\v1\\<^sup>2 = 2" by (metis complex_cnj_one cpx_vec_length_inner_prod mult.left_neutral of_real_eq_iff of_real_numeral one_add_one) then have "\U * |v1\\\<^sup>2 = 2" using c0 d1 assms(2) unit_vec_to_col by simp moreover have "col U i + 1 \\<^sub>v col U j = U * |v1\" using c0 d1 a0 sum_of_unit_vec_to_col by blast moreover have "col U i + 1 \\<^sub>v col U j = col U i + col U j" by simp ultimately have "\col U i + col U j|col U i + col U j\ = 2" using cpx_vec_length_inner_prod by (metis of_real_numeral) moreover have "\col U i + col U j|col U i + col U j\ = \col U i|col U i\ + \col U j|col U i\ + \col U i|col U j\ + \col U j|col U j\" using inner_prod_is_sesquilinear[of "col U i" "dim_row U" "col U j" "col U i" "col U j" "1" "1" "1" "1"] by simp ultimately have f2:"\col U j|col U i\ + \col U i|col U j\ = 0" using c0 a0 f1 inner_prod_def lessThan_atLeast0 by simp have "\v2\\<^sup>2 = 1 + cnj \ * \" using a0 a2 d2 sum_of_unit_vec_length by simp then have "\v2\\<^sup>2 = 2" by (metis Re_complex_of_real complex_norm_square mult.commute norm_ii numeral_Bit0 numeral_One numeral_eq_one_iff of_real_numeral one_power2) moreover have "\U * |v2\\\<^sup>2 = \v2\\<^sup>2" using c0 d2 assms(2) unit_vec_to_col by simp moreover have "\col U i + \ \\<^sub>v col U j|col U i + \ \\<^sub>v col U j\ = \U * |v2\\\<^sup>2" using c0 a0 d2 sum_of_unit_vec_to_col cpx_vec_length_inner_prod by auto moreover have "\col U i + \ \\<^sub>v col U j|col U i + \ \\<^sub>v col U j\ = \col U i|col U i\ + (-\) * \col U j|col U i\ + \ * \col U i|col U j\ + \col U j|col U j\" using inner_prod_is_sesquilinear[of "col U i" "dim_row U" "col U j" "col U i" "col U j" "1" "\" "1" "\"] by simp ultimately have "\col U j|col U i\ - \col U i|col U j\ = 0" using c0 a0 f1 inner_prod_def lessThan_atLeast0 by auto then show "(\k\ * U) $$ (i, j) = 1\<^sub>m (dim_col U) $$ (i, j)" using c0 assms(1) a0 one_mat_def dagger_def by auto qed then have "(U\<^sup>\) * U = 1\<^sub>m n" using c0 by simp then have "inverts_mat (U\<^sup>\) U" using c1 inverts_mat_def by auto then have "inverts_mat U (U\<^sup>\)" using c0 c1 inverts_mat_sym by simp then have "U * (U\<^sup>\) = 1\<^sub>m (dim_row U)" using c0 inverts_mat_def by auto then show ?thesis using f0 unitary_def by simp qed lemma inner_prod_with_unitary_mat [simp]: assumes "unitary U" and "dim_vec u = dim_col U" and "dim_vec v = dim_col U" shows "\U * |u\|U * |v\\ = \u|v\" proof - have f1:"\U * |u\|U * |v\\ = (\|u\| * (U\<^sup>\) * U * |v\) $$ (0,0)" using assms(2-3) bra_mat_on_vec mult_ket_vec_is_ket_vec_of_mult - by (smt assoc_mult_mat carrier_mat_triv col_fst_def dim_vec dim_col_of_dagger index_mult_mat(2) + by (smt (verit, ccfv_threshold) assoc_mult_mat carrier_mat_triv col_fst_def dim_vec dim_col_of_dagger index_mult_mat(2) index_mult_mat(3) inner_prod_with_times_mat ket_vec_def mat_carrier) moreover have f2:"\|u\| \ carrier_mat 1 (dim_vec v)" using bra_def ket_vec_def assms(2-3) by simp moreover have f3:"U\<^sup>\ \ carrier_mat (dim_col U) (dim_row U)" using dagger_def by simp ultimately have "\U * |u\|U * |v\\ = (\|u\| * (U\<^sup>\ * U) * |v\) $$ (0,0)" using assms(3) assoc_mult_mat by (metis carrier_mat_triv) also have "\ = (\|u\| * |v\) $$ (0,0)" using assms(1) unitary_def by (simp add: assms(2) bra_def ket_vec_def) finally show ?thesis using assms(2-3) inner_prod_with_times_mat by presburger qed text \As a consequence we prove that columns and rows of a unitary matrix are orthonormal vectors.\ lemma unitary_unit_col [simp]: assumes "unitary U" and "dim_col U = n" and "i < n" shows "\col U i\ = 1" using assms unit_vec_to_col unitary_is_length_preserving by simp lemma unitary_unit_row [simp]: assumes "unitary U" and "dim_row U = n" and "i < n" shows "\row U i\ = 1" proof - have "row U i = col (U\<^sup>t) i" using assms(2-3) by simp thus ?thesis using assms transpose_of_unitary_is_unitary unitary_unit_col by (metis index_transpose_mat(3)) qed lemma orthogonal_col_of_unitary [simp]: assumes "unitary U" and "dim_col U = n" and "i < n" and "j < n" and "i \ j" shows "\col U i|col U j\ = 0" proof - have "\col U i|col U j\ = \U * |unit_vec n i\| U * |unit_vec n j\\" using assms(2-4) unit_vec_to_col by simp also have "\ = \unit_vec n i |unit_vec n j\" using assms(1-2) inner_prod_with_unitary_mat index_unit_vec(3) by simp finally show ?thesis using assms(3-5) by simp qed lemma orthogonal_row_of_unitary [simp]: fixes U::"complex mat" assumes "unitary U" and "dim_row U = n" and "i < n" and "j < n" and "i \ j" shows "\row U i|row U j\ = 0" using assms orthogonal_col_of_unitary transpose_of_unitary_is_unitary col_transpose by (metis index_transpose_mat(3)) text\ As a consequence, we prove that a quantum gate acting on a state of a system of n qubits give another state of that same system. \ lemma gate_on_state_is_state [intro, simp]: assumes a1:"gate n A" and a2:"state n v" shows "state n (A * v)" proof show "dim_row (A * v) = 2^n" using gate_def state_def a1 by simp next show "dim_col (A * v) = 1" using state_def a2 by simp next have "square_mat A" using a1 gate_def by simp then have "dim_col A = 2^n" using a1 gate.dim_row by simp then have "dim_col A = dim_row v" using a2 state.dim_row by simp then have "\col (A * v) 0\ = \col v 0\" using unitary_is_length_preserving_bis assms gate_def state_def by simp thus"\col (A * v) 0\ = 1" using a2 state.is_normal by simp qed subsection \A Few Well-known Quantum Gates\ text \ Any unitary operation on n qubits can be implemented exactly by composing single qubits and CNOT-gates (controlled-NOT gates). However, no straightforward method is known to implement these gates in a fashion which is resistant to errors. But, the Hadamard gate, the phase gate, the CNOT-gate and the @{text "\/8"} gate are also universal for quantum computations, i.e. any quantum circuit on n qubits can be approximated to an arbitrary accuracy by using only these gates, and these gates can be implemented in a fault-tolerant way. \ text \We introduce a coercion from real matrices to complex matrices.\ definition real_to_cpx_mat:: "real mat \ complex mat" where "real_to_cpx_mat A \ mat (dim_row A) (dim_col A) (\(i,j). A $$ (i,j))" text \Our first quantum gate: the identity matrix! Arguably, not a very interesting one though!\ definition Id :: "nat \ complex mat" where "Id n \ 1\<^sub>m (2^n)" lemma id_is_gate [simp]: "gate n (Id n)" proof show "dim_row (Id n) = 2^n" using Id_def by simp next show "square_mat (Id n)" using Id_def by simp next show "unitary (Id n)" by (simp add: Id_def) qed text \More interesting: the Pauli matrices.\ definition X ::"complex mat" where "X \ mat 2 2 (\(i,j). if i=j then 0 else 1)" text\ Be aware that @{text "gate n A"} means that the matrix A has dimension @{text "2^n * 2^n"}. For instance, with this convention a 2 X 2 matrix A which is unitary satisfies @{text "gate 1 A"} but not @{text "gate 2 A"} as one might have been expected. \ lemma dagger_of_X [simp]: "X\<^sup>\ = X" using dagger_def by (simp add: X_def cong_mat) lemma X_inv [simp]: "X * X = 1\<^sub>m 2" apply(simp add: X_def times_mat_def one_mat_def) apply(rule cong_mat) by(auto simp: scalar_prod_def) lemma X_is_gate [simp]: "gate 1 X" by (simp add: gate_def unitary_def) (simp add: X_def) definition Y ::"complex mat" where "Y \ mat 2 2 (\(i,j). if i=j then 0 else (if i=0 then -\ else \))" lemma dagger_of_Y [simp]: "Y\<^sup>\ = Y" using dagger_def by (simp add: Y_def cong_mat) lemma Y_inv [simp]: "Y * Y = 1\<^sub>m 2" apply(simp add: Y_def times_mat_def one_mat_def) apply(rule cong_mat) by(auto simp: scalar_prod_def) lemma Y_is_gate [simp]: "gate 1 Y" by (simp add: gate_def unitary_def) (simp add: Y_def) definition Z ::"complex mat" where "Z \ mat 2 2 (\(i,j). if i\j then 0 else (if i=0 then 1 else -1))" lemma dagger_of_Z [simp]: "Z\<^sup>\ = Z" using dagger_def by (simp add: Z_def cong_mat) lemma Z_inv [simp]: "Z * Z = 1\<^sub>m 2" apply(simp add: Z_def times_mat_def one_mat_def) apply(rule cong_mat) by(auto simp: scalar_prod_def) lemma Z_is_gate [simp]: "gate 1 Z" by (simp add: gate_def unitary_def) (simp add: Z_def) text \The Hadamard gate\ definition H ::"complex mat" where "H \ 1/sqrt(2) \\<^sub>m (mat 2 2 (\(i,j). if i\j then 1 else (if i=0 then 1 else -1)))" lemma H_without_scalar_prod: "H = mat 2 2 (\(i,j). if i\j then 1/sqrt(2) else (if i=0 then 1/sqrt(2) else -(1/sqrt(2))))" using cong_mat by (auto simp: H_def) lemma dagger_of_H [simp]: "H\<^sup>\ = H" using dagger_def by (auto simp: H_def cong_mat) lemma H_inv [simp]: "H * H = 1\<^sub>m 2" apply(simp add: H_def times_mat_def one_mat_def) apply(rule cong_mat) by(auto simp: scalar_prod_def complex_eqI) lemma H_is_gate [simp]: "gate 1 H" by (simp add: gate_def unitary_def) (simp add: H_def) lemma H_values: fixes i j:: nat assumes "i < dim_row H" and "j < dim_col H" and "i \ 1 \ j \ 1" shows "H $$ (i,j) = 1/sqrt 2" proof- have "i < 2" using assms(1) by (simp add: H_without_scalar_prod less_2_cases) moreover have "j < 2" using assms(2) by (simp add: H_without_scalar_prod less_2_cases) ultimately show ?thesis - using assms(3) H_without_scalar_prod by(smt One_nat_def index_mat(1) less_2_cases old.prod.case) + using assms(3) H_without_scalar_prod by (smt (verit) One_nat_def index_mat(1) less_2_cases old.prod.case) qed lemma H_values_right_bottom: fixes i j:: nat assumes "i = 1 \ j = 1" shows "H $$ (i,j) = - 1/sqrt 2" using assms by (simp add: H_without_scalar_prod) text \The controlled-NOT gate\ definition CNOT ::"complex mat" where "CNOT \ mat 4 4 (\(i,j). if i=0 \ j=0 then 1 else (if i=1 \ j=1 then 1 else (if i=2 \ j=3 then 1 else (if i=3 \ j=2 then 1 else 0))))" lemma dagger_of_CNOT [simp]: "CNOT\<^sup>\ = CNOT" using dagger_def by (simp add: CNOT_def cong_mat) lemma CNOT_inv [simp]: "CNOT * CNOT = 1\<^sub>m 4" apply(simp add: CNOT_def times_mat_def one_mat_def) apply(rule cong_mat) by(auto simp: scalar_prod_def) lemma CNOT_is_gate [simp]: "gate 2 CNOT" by (simp add: gate_def unitary_def) (simp add: CNOT_def) text \The phase gate, also known as the S-gate\ definition S ::"complex mat" where "S \ mat 2 2 (\(i,j). if i=0 \ j=0 then 1 else (if i=1 \ j=1 then \ else 0))" text \The @{text "\/8"} gate, also known as the T-gate\ definition T ::"complex mat" where "T \ mat 2 2 (\(i,j). if i=0 \ j=0 then 1 else (if i=1 \ j=1 then exp(\*(pi/4)) else 0))" text \A few relations between the Hadamard gate and the Pauli matrices\ lemma HXH_is_Z [simp]: "H * X * H = Z" apply(simp add: X_def Z_def H_def times_mat_def) apply(rule cong_mat) by(auto simp add: scalar_prod_def complex_eqI) lemma HYH_is_minusY [simp]: "H * Y * H = - Y" apply(simp add: Y_def H_def times_mat_def) apply(rule eq_matI) by(auto simp add: scalar_prod_def complex_eqI) lemma HZH_is_X [simp]: shows "H * Z * H = X" apply(simp add: X_def Z_def H_def times_mat_def) apply(rule cong_mat) by(auto simp add: scalar_prod_def complex_eqI) subsection \The Bell States\ text \ We introduce below the so-called Bell states, also known as EPR pairs (EPR stands for Einstein, Podolsky and Rosen). \ definition bell00 ::"complex mat" ("|\\<^sub>0\<^sub>0\") where "bell00 \ 1/sqrt(2) \\<^sub>m |vec 4 (\i. if i=0 \ i=3 then 1 else 0)\" definition bell01 ::"complex mat" ("|\\<^sub>0\<^sub>1\") where "bell01 \ 1/sqrt(2) \\<^sub>m |vec 4 (\i. if i=1 \ i=2 then 1 else 0)\" definition bell10 ::"complex mat" ("|\\<^sub>1\<^sub>0\") where "bell10 \ 1/sqrt(2) \\<^sub>m |vec 4 (\i. if i=0 then 1 else if i=3 then -1 else 0)\" definition bell11 ::"complex mat" ("|\\<^sub>1\<^sub>1\") where "bell11 \ 1/sqrt(2) \\<^sub>m |vec 4 (\i. if i=1 then 1 else if i=2 then -1 else 0)\" lemma shows bell00_is_state [simp]:"state 2 |\\<^sub>0\<^sub>0\" and bell01_is_state [simp]:"state 2 |\\<^sub>0\<^sub>1\" and bell10_is_state [simp]:"state 2 |\\<^sub>1\<^sub>0\" and bell11_is_state [simp]:"state 2 |\\<^sub>1\<^sub>1\" by (auto simp: state_def bell00_def bell01_def bell10_def bell11_def ket_vec_def) (auto simp: cpx_vec_length_def Set_Interval.lessThan_atLeast0 cmod_def power2_eq_square) lemma bell00_index [simp]: shows "|\\<^sub>0\<^sub>0\ $$ (0,0) = 1/sqrt 2" and "|\\<^sub>0\<^sub>0\ $$ (1,0) = 0" and "|\\<^sub>0\<^sub>0\ $$ (2,0) = 0" and "|\\<^sub>0\<^sub>0\ $$ (3,0) = 1/sqrt 2" by (auto simp: bell00_def ket_vec_def) lemma bell01_index [simp]: shows "|\\<^sub>0\<^sub>1\ $$ (0,0) = 0" and "|\\<^sub>0\<^sub>1\ $$ (1,0) = 1/sqrt 2" and "|\\<^sub>0\<^sub>1\ $$ (2,0) = 1/sqrt 2" and "|\\<^sub>0\<^sub>1\ $$ (3,0) = 0" by (auto simp: bell01_def ket_vec_def) lemma bell10_index [simp]: shows "|\\<^sub>1\<^sub>0\ $$ (0,0) = 1/sqrt 2" and "|\\<^sub>1\<^sub>0\ $$ (1,0) = 0" and "|\\<^sub>1\<^sub>0\ $$ (2,0) = 0" and "|\\<^sub>1\<^sub>0\ $$ (3,0) = - 1/sqrt 2" by (auto simp: bell10_def ket_vec_def) lemma bell_11_index [simp]: shows "|\\<^sub>1\<^sub>1\ $$ (0,0) = 0" and "|\\<^sub>1\<^sub>1\ $$ (1,0) = 1/sqrt 2" and "|\\<^sub>1\<^sub>1\ $$ (2,0) = - 1/sqrt 2" and "|\\<^sub>1\<^sub>1\ $$ (3,0) = 0" by (auto simp: bell11_def ket_vec_def) subsection \The Bitwise Inner Product\ definition bitwise_inner_prod:: "nat \ nat \ nat \ nat" where "bitwise_inner_prod n i j = (\k\{0.. nat \ nat \ nat" ("_ \\<^bsub>_\<^esub> _") where "bip i n j \ bitwise_inner_prod n i j" lemma bitwise_inner_prod_fst_el_0: assumes "i < 2^n \ j < 2^n" shows "(i \\<^bsub>Suc n\<^esub> j) = (i mod 2^n) \\<^bsub>n\<^esub> (j mod 2^n)" proof- have "bip i (Suc n) j = (\k\{0..<(Suc n)}. (bin_rep (Suc n) i) ! k * (bin_rep (Suc n) j) ! k)" using bitwise_inner_prod_def by simp also have "... = bin_rep (Suc n) i ! 0 * bin_rep (Suc n) j ! 0 + (\k\{1..<(Suc n)}. bin_rep (Suc n) i ! k * bin_rep (Suc n) j ! k)" by (simp add: sum.atLeast_Suc_lessThan) also have "... = (\k\{1..<(Suc n)}. bin_rep (Suc n) i ! k * bin_rep (Suc n) j ! k)" using bin_rep_index_0[of i n] bin_rep_index_0[of j n] assms by auto also have "... = (\k\{0..k. bin_rep (Suc n) i ! k * bin_rep (Suc n) j ! k" "0" "n"] by (metis (no_types, lifting) One_nat_def add.commute plus_1_eq_Suc sum.cong) finally have "bip i (Suc n) j = (\k\{0..{0..n} \ bin_rep (Suc n) i ! (k+1) = bin_rep n (i mod 2^n) ! k" for k using bin_rep_def by (simp add: bin_rep_aux_neq_nil) moreover have "k\{0..n} \ bin_rep (Suc n) j !(k+1) = bin_rep n (j mod 2^n) ! k" for k using bin_rep_def by (simp add: bin_rep_aux_neq_nil) ultimately show ?thesis using assms bin_rep_index_0 bitwise_inner_prod_def by simp qed lemma bitwise_inner_prod_fst_el_is_1: fixes n i j:: nat assumes "i \ 2^n \ j \ 2^n" and "i < 2^(n+1) \ j < 2^(n+1)" shows "(i \\<^bsub>(n+1)\<^esub> j) = 1 + ((i mod 2^n) \\<^bsub>n\<^esub> (j mod 2^n))" proof- have "bip i (Suc n) j = (\k\{0..<(Suc n)}. bin_rep (Suc n) i ! k * bin_rep (Suc n) j ! k)" using bitwise_inner_prod_def by simp also have "... = bin_rep (Suc n) i ! 0 * bin_rep (Suc n) j ! 0 + (\k\{1..<(Suc n)}. bin_rep (Suc n) i ! k * bin_rep (Suc n) j ! k)" by (simp add: sum.atLeast_Suc_lessThan) also have "... = 1 + (\k\{1..<(Suc n)}. bin_rep (Suc n) i ! k * bin_rep (Suc n) j ! k)" using bin_rep_index_0_geq[of n i] bin_rep_index_0_geq[of n j] assms by simp also have "... = 1 + (\k \ {0..k. (bin_rep (Suc n) i)!k * (bin_rep (Suc n) j)!k" "0" "n"] by (metis (no_types, lifting) One_nat_def Suc_eq_plus1 sum.cong) finally have f0:"bip i (Suc n) j = 1 + (\k\{0..{0..n} \ bin_rep (Suc n) i ! (k+1) = bin_rep n (i mod 2^n) ! k \ bin_rep (Suc n) j ! (k+1) = bin_rep n (j mod 2^n) ! k" for k using bin_rep_def by(metis Suc_eq_plus1 bin_rep_aux.simps(2) bin_rep_aux_neq_nil butlast.simps(2) nth_Cons_Suc) ultimately show ?thesis using bitwise_inner_prod_def by simp qed lemma bitwise_inner_prod_with_zero: assumes "m < 2^n" shows "(0 \\<^bsub>n\<^esub> m) = 0" proof- have "(0 \\<^bsub>n\<^esub> m) = (\j\{0..j\{0..j\{0..Quantum Teleportation\ theory Quantum_Teleportation imports More_Tensor Basics Measurement begin definition alice:: "complex Matrix.mat \ complex Matrix.mat" where "alice \ \ (H \ Id 2) * ((CNOT \ Id 1) * (\ \ |\\<^sub>0\<^sub>0\))" abbreviation M1:: "complex Matrix.mat" where "M1 \ mat_of_cols_list 8 [[1, 0, 0, 0, 0, 0, 0, 0], [0, 1, 0, 0, 0, 0, 0, 0], [0, 0, 1, 0, 0, 0, 0, 0], [0, 0, 0, 1, 0, 0, 0, 0], [0, 0, 0, 0, 0, 0, 1, 0], [0, 0, 0, 0, 0, 0, 0, 1], [0, 0, 0, 0, 1, 0, 0, 0], [0, 0, 0, 0, 0, 1, 0, 0]]" lemma tensor_prod_of_cnot_id_1: shows "(CNOT \ Id 1) = M1" proof show "dim_col (CNOT \ Id 1) = dim_col M1" by(simp add: CNOT_def Id_def mat_of_cols_list_def) show "dim_row (CNOT \ Id 1) = dim_row M1" by(simp add: CNOT_def Id_def mat_of_cols_list_def) fix i j::nat assume "i < dim_row M1" and "j < dim_col M1" then have "i \ {0..<8} \ j \ {0..<8}" by (auto simp add: mat_of_cols_list_def) then show "(CNOT \ Id 1) $$ (i, j) = M1 $$ (i, j)" by (auto simp add: Id_def CNOT_def mat_of_cols_list_def) qed abbreviation M2:: "complex Matrix.mat" where "M2 \ mat_of_cols_list 8 [[1/sqrt(2), 0, 0, 0, 1/sqrt(2), 0, 0, 0], [0, 1/sqrt(2), 0, 0, 0, 1/sqrt(2), 0, 0], [0, 0, 1/sqrt(2), 0, 0, 0, 1/sqrt(2), 0], [0, 0, 0, 1/sqrt(2), 0, 0, 0, 1/sqrt(2)], [1/sqrt(2), 0, 0, 0, -1/sqrt(2), 0, 0, 0], [0, 1/sqrt(2), 0, 0, 0, -1/sqrt(2), 0, 0], [0, 0, 1/sqrt(2), 0, 0, 0, -1/sqrt(2), 0], [0, 0, 0, 1/sqrt(2), 0, 0, 0, -1/sqrt(2)]]" lemma tensor_prod_of_h_id_2: shows "(H \ Id 2) = M2" proof show "dim_col (H \ Id 2) = dim_col M2" by(simp add: H_def Id_def mat_of_cols_list_def) show "dim_row (H \ Id 2) = dim_row M2" by(simp add: H_def Id_def mat_of_cols_list_def) fix i j::nat assume "i < dim_row M2" and "j < dim_col M2" then have "i \ {0..<8} \ j \ {0..<8}" by (auto simp add: mat_of_cols_list_def) then show "(H \ Id 2) $$ (i, j) = M2 $$ (i, j)" by (auto simp add: Id_def H_def mat_of_cols_list_def) qed lemma alice_step_1_state [simp]: assumes "state 1 \" shows "state 3 (\ \ |\\<^sub>0\<^sub>0\)" using assms bell00_is_state tensor_state by(metis One_nat_def Suc_1 numeral_3_eq_3 plus_1_eq_Suc) lemma alice_step_2_state: assumes "state 1 \" shows "state 3 ((CNOT \ Id 1) * (\ \ |\\<^sub>0\<^sub>0\))" proof- have "gate 3 (CNOT \ Id 1)" using CNOT_is_gate id_is_gate tensor_gate by (metis numeral_plus_one semiring_norm(5)) then show "state 3 ((CNOT \ Id 1) * (\ \ |\\<^sub>0\<^sub>0\))" using assms by simp qed lemma alice_state [simp]: assumes "state 1 \" shows "state 3 (alice \) " proof- have "gate 3 (H \ Id 2)" using tensor_gate id_is_gate H_is_gate by(metis eval_nat_numeral(3) plus_1_eq_Suc) then show ?thesis using assms alice_step_2_state by(simp add: alice_def) qed lemma alice_step_1: assumes "state 1 \" and "\ = \ $$ (0,0)" and "\ = \ $$ (1,0)" shows "(\ \ |\\<^sub>0\<^sub>0\) = mat_of_cols_list 8 [[\/sqrt(2),0,0,\/sqrt(2),\/sqrt(2),0,0,\/sqrt(2)]]" proof define v where asm:"v = mat_of_cols_list 8 [[\/sqrt(2),0,0,\/sqrt(2),\/sqrt(2),0,0,\/sqrt(2)]]" then show "dim_row (\ \ |\\<^sub>0\<^sub>0\) = dim_row v" using assms(1) alice_step_1_state state.dim_row mat_of_cols_list_def by fastforce show "dim_col (\ \ |\\<^sub>0\<^sub>0\) = dim_col v" using assms(1) alice_step_1_state state.is_column asm mat_of_cols_list_def by fastforce show "\i j. i < dim_row v \ j < dim_col v \ (\ \ |\\<^sub>0\<^sub>0\) $$ (i,j) = v $$ (i,j)" proof- fix i j assume "i < dim_row v" and "j < dim_col v" then have "i \ {0..<8} \ j = 0" using asm by (auto simp add: mat_of_cols_list_def) moreover have "dim_row |\\<^sub>0\<^sub>0\ = 4" using bell00_is_state state_def by simp moreover have "dim_col |\\<^sub>0\<^sub>0\ = 1" using bell00_is_state state_def by simp ultimately show "(\ \ |\\<^sub>0\<^sub>0\) $$ (i, j) = v $$ (i,j)" using state_def assms asm by (auto simp add: bell00_def) qed qed lemma alice_step_2: assumes "state 1 \" and "\ = \ $$ (0,0)" and "\ = \ $$ (1,0)" shows "(CNOT \ Id 1) * (\ \ |\\<^sub>0\<^sub>0\) = mat_of_cols_list 8 [[\/sqrt(2),0,0,\/sqrt(2),0,\/sqrt(2),\/sqrt(2),0]]" proof have f0:"(\ \ |\\<^sub>0\<^sub>0\) = mat_of_cols_list 8 [[\/sqrt(2),0,0,\/sqrt(2),\/sqrt(2),0,0,\/sqrt(2)]]" using assms alice_step_1 by simp define v where asm:"v = mat_of_cols_list 8 [[\/sqrt(2),0,0,\/sqrt(2),0,\/sqrt(2),\/sqrt(2),0]]" then show "dim_row ((CNOT \ Id 1) * (\ \ |\\<^sub>0\<^sub>0\)) = dim_row v" using assms(1) alice_step_2_state state.dim_row mat_of_cols_list_def by fastforce show "dim_col ((CNOT \ Id 1) * (\ \ |\\<^sub>0\<^sub>0\)) = dim_col v" using assms(1) alice_step_2_state state.is_column asm mat_of_cols_list_def by fastforce show "\i j. i j ((CNOT \ Id 1) * (\ \ |\\<^sub>0\<^sub>0\)) $$ (i,j) = v $$ (i,j)" proof- fix i j assume "i < dim_row v" and "j < dim_col v" then have "i \ {0..<8::nat} \ j = 0" using asm by (auto simp add: mat_of_cols_list_def) then have "(M1 * (\ \ |\\<^sub>0\<^sub>0\)) $$ (i,j) = v $$ (i,j)" by (auto simp add: f0 asm mat_of_cols_list_def times_mat_def scalar_prod_def) then show "((CNOT \ Id 1) * (\ \ |\\<^sub>0\<^sub>0\)) $$ (i,j) = v $$ (i,j)" using tensor_prod_of_cnot_id_1 by simp qed qed lemma alice_result: assumes "state 1 \" and "\ = \ $$ (0,0)" and "\ = \ $$ (1,0)" shows "alice \ = mat_of_cols_list 8 [[\/2, \/2, \/2, \/2, \/2, -\/2, -\/2, \/2]]" proof define v where a0:"v = mat_of_cols_list 8 [[\/2, \/2, \/2, \/2, \/2, -\/2, -\/2, \/2]]" define w where a1:"w = (CNOT \ Id 1) * (\ \ |\\<^sub>0\<^sub>0\)" then have f0:"w = mat_of_cols_list 8 [[\/sqrt(2), 0, 0, \/sqrt(2), 0, \/sqrt(2), \/sqrt(2), 0]]" using assms alice_step_2 by simp show "dim_row (alice \) = dim_row v" using assms(1) alice_state state_def a0 by (simp add: mat_of_cols_list_def) show "dim_col (alice \) = dim_col v" using assms(1) alice_state state_def a0 by (simp add: mat_of_cols_list_def) show "\i j. i < dim_row v \ j < dim_col v \ alice \ $$ (i,j) = v $$ (i,j)" proof- fix i j assume "i < dim_row v" and "j < dim_col v" then have "i \ {0..<8} \ j = 0" using a0 by (auto simp add: Tensor.mat_of_cols_list_def) then have "(M2 * w) $$ (i,j) = v $$ (i,j)" by (auto simp add: f0 a0 mat_of_cols_list_def times_mat_def scalar_prod_def) then show "alice \ $$ (i,j) = v $$ (i,j)" by (simp add: tensor_prod_of_h_id_2 alice_def a1) qed qed text \ An application of function @{term "alice"} to a state @{term "\"} of a 1-qubit system results in the following cases. \ definition alice_meas:: "complex Matrix.mat \ _list" where "alice_meas \ = [ ((prob0 3 (alice \) 0) * (prob0 3 (post_meas0 3 (alice \) 0) 1), post_meas0 3 (post_meas0 3 (alice \) 0) 1) , ((prob0 3 (alice \) 0) * (prob1 3 (post_meas0 3 (alice \) 0) 1), post_meas1 3 (post_meas0 3 (alice \) 0) 1) , ((prob1 3 (alice \) 0) * (prob0 3 (post_meas1 3 (alice \) 0) 1), post_meas0 3 (post_meas1 3 (alice \) 0) 1) , ((prob1 3 (alice \) 0) * (prob1 3 (post_meas1 3 (alice \) 0) 1), post_meas1 3 (post_meas1 3 (alice \) 0) 1) ]" definition alice_pos:: "complex Matrix.mat \ complex Matrix.mat \ bool" where "alice_pos \ q \ q = mat_of_cols_list 8 [[\ $$ (0,0), \ $$ (1,0), 0, 0, 0, 0, 0, 0]] \ q = mat_of_cols_list 8 [[0, 0, \ $$ (1,0), \ $$ (0,0), 0, 0, 0, 0]] \ q = mat_of_cols_list 8 [[0, 0, 0, 0, \ $$ (0,0), - \ $$ (1,0), 0, 0]] \ q = mat_of_cols_list 8 [[0, 0, 0, 0, 0, 0, - \ $$ (1,0), \ $$ (0,0)]]" lemma phi_vec_length: assumes "state 1 \" shows "cmod(\ $$ (0,0))^2 + cmod(\ $$ (Suc 0,0))^2 = 1" using set_2 assms state_def Matrix.col_def cpx_vec_length_def by(auto simp add: atLeast0LessThan) lemma select_index_3_subsets [simp]: shows "{j::nat. select_index 3 0 j} = {4,5,6,7} \ {j::nat. j < 8 \ \ select_index 3 0 j} = {0,1,2,3} \ {j::nat. select_index 3 1 j} = {2,3,6,7} \ {j::nat. j < 8 \ \ select_index 3 1 j} = {0,1,4,5}" proof- have "{j::nat. select_index 3 0 j} = {4,5,6,7}" by (auto simp add: select_index_def) moreover have "{j::nat. j < 8 \ \ select_index 3 0 j} = {0,1,2,3}" by(auto simp add: select_index_def) moreover have f1:"{j::nat. select_index 3 1 j} = {2,3,6,7}" proof show "{j. select_index 3 1 j} \ {2,3,6,7}" proof fix j::nat assume "j \ {j. select_index 3 1 j}" then have "j \ {0..<8} \ j mod 4 \ {2,3}" by (auto simp add: select_index_def) then show "j \ {2,3,6,7}" by auto qed show "{2,3,6,7} \ {j. select_index 3 1 j}" by (auto simp add: select_index_def) qed moreover have "{j::nat. j < 8 \ \ select_index 3 1 j} = {0,1,4,5}" proof- have "{j::nat. j < 8 \ j \ {2,3,6,7}} = {0,1,4,5}" by auto then show ?thesis using f1 by blast qed ultimately show ?thesis by simp qed lemma prob_index_0_alice: assumes "state 1 \" shows "prob0 3 (alice \) 0 = 1/2 \ prob1 3 (alice \) 0 = 1/2" proof show "prob0 3 (alice \) 0 = 1/2" using alice_result assms prob0_def alice_state apply auto by (metis (no_types, opaque_lifting) One_nat_def phi_vec_length four_x_squared mult.commute nonzero_mult_div_cancel_right times_divide_eq_right zero_neq_numeral) then show"prob1 3 (alice \) 0 = 1/2" using prob_sum_is_one[of "3" "alice \" "0"] alice_state[of "\"] assms by linarith qed lemma post_meas0_index_0_alice: assumes "state 1 \" and "\ = \ $$ (0,0)" and "\ = \ $$ (1,0)" shows "post_meas0 3 (alice \) 0 = mat_of_cols_list 8 [[\/sqrt(2), \/sqrt(2), \/sqrt(2), \/sqrt(2), 0, 0, 0, 0]]" proof define v where asm:"v = mat_of_cols_list 8 [[\/sqrt(2),\/sqrt(2),\/sqrt(2),\/sqrt(2),0,0,0,0]]" then show "dim_row (post_meas0 3 (alice \) 0) = dim_row v" using mat_of_cols_list_def post_meas0_def assms(1) alice_state ket_vec_def by simp show "dim_col (post_meas0 3 (alice \) 0) = dim_col v" using mat_of_cols_list_def post_meas0_def assms(1) alice_state ket_vec_def asm by simp show "\i j. i j post_meas0 3 (alice \) 0 $$ (i,j) = v $$ (i,j)" proof- fix i j assume "i < dim_row v" and "j < dim_col v" then have "i \ {0..<8} \ j = 0" using asm set_8_atLeast0 mat_of_cols_list_def by auto then show "post_meas0 3 (alice \) 0 $$ (i, j) = v $$ (i, j)" using post_meas0_def assms asm mat_of_cols_list_def ket_vec_def apply (auto simp add: prob_index_0_alice) using assms(1) alice_result select_index_def by auto qed qed lemma post_meas1_index_0_alice: assumes "state 1 \" and "\ = \ $$ (0,0)" and "\ = \ $$ (1,0)" shows "post_meas1 3 (alice \) 0 = mat_of_cols_list 8 [[0,0,0,0,\/sqrt(2),-\/sqrt(2),-\/sqrt(2),\/sqrt(2)]]" proof define v where asm:"v = mat_of_cols_list 8 [[0,0,0,0,\/sqrt(2),-\/sqrt(2),-\/sqrt(2),\/sqrt(2)]]" then show "dim_row (post_meas1 3 (alice \) 0) = dim_row v" using mat_of_cols_list_def post_meas1_def assms(1) alice_state ket_vec_def by simp show "dim_col (post_meas1 3 (alice \) 0) = dim_col v" using mat_of_cols_list_def post_meas1_def assms(1) alice_state ket_vec_def asm by simp show "\i j. i j post_meas1 3 (alice \) 0 $$ (i,j) = v $$ (i,j)" proof- fix i j assume "i < dim_row v" and "j < dim_col v" then have "i \ {0..<8} \ j = 0" using asm set_8_atLeast0 mat_of_cols_list_def by auto then show "post_meas1 3 (alice \) 0 $$ (i,j) = v $$ (i,j)" using post_meas1_def assms asm mat_of_cols_list_def ket_vec_def apply (auto simp add: prob_index_0_alice) using assms(1) alice_result select_index_def by auto qed qed lemma post_meas0_index_0_alice_state [simp]: assumes "state 1 \" shows "state 3 (post_meas0 3 (alice \) 0)" using assms by (simp add: prob_index_0_alice) lemma post_meas1_index_0_alice_state [simp]: assumes "state 1 \" shows "state 3 (post_meas1 3 (alice \) 0)" using assms by (simp add: prob_index_0_alice) lemma Alice_case [simp]: assumes "state 1 \" and "state 3 q" and "List.member (alice_meas \) (p, q)" shows "alice_pos \ q" proof- define \ \ where a0:"\ = \ $$ (0,0)" and a1:"\ = \ $$ (1,0)" have f0:"prob0 3 (Matrix.mat 8 (Suc 0) (\(i,j). [[\ $$ (0,0)/sqrt 2, \ $$ (Suc 0,0)/sqrt 2, \ $$ (Suc 0,0)/sqrt 2, \ $$ (0,0)/sqrt 2,0,0,0,0]]!j!i)) (Suc 0) = 1/2" using post_meas0_index_0_alice prob0_def mat_of_cols_list_def post_meas0_index_0_alice_state assms(1) a0 a1 select_index_3_subsets by (auto simp add: norm_divide power_divide phi_vec_length) have f1:"prob1 3 (Matrix.mat 8 (Suc 0) (\(i,j). [[\ $$ (0,0)/sqrt 2, \ $$ (Suc 0,0)/sqrt 2, \ $$ (Suc 0,0)/sqrt 2, \ $$ (0,0)/sqrt 2, 0, 0, 0, 0]] ! j ! i)) (Suc 0) = 1/2" using post_meas0_index_0_alice prob1_def mat_of_cols_list_def post_meas0_index_0_alice_state assms(1) a0 a1 select_index_3_subsets by(auto simp add: norm_divide power_divide phi_vec_length algebra_simps) have f2:"prob0 3 (Matrix.mat 8 (Suc 0) (\(i,j). [[0,0,0,0,\ $$ (0,0)/complex_of_real (sqrt 2),-(\ $$ (Suc 0,0)/complex_of_real (sqrt 2)), -(\ $$ (Suc 0,0)/complex_of_real (sqrt 2)),\ $$ (0,0)/complex_of_real (sqrt 2)]] ! j ! i)) (Suc 0) = 1/2" using post_meas1_index_0_alice prob0_def mat_of_cols_list_def post_meas1_index_0_alice_state assms(1) a0 a1 select_index_3_subsets by(auto simp add: norm_divide power_divide phi_vec_length) have f3:"prob1 3 (Matrix.mat 8 (Suc 0) (\(i,j). [[0,0,0,0,\ $$ (0,0)/complex_of_real (sqrt 2),-(\ $$ (Suc 0,0)/complex_of_real (sqrt 2)), -(\ $$ (Suc 0,0)/complex_of_real (sqrt 2)), \ $$ (0,0)/complex_of_real (sqrt 2)]] ! j ! i)) (Suc 0) = 1/2" using post_meas1_index_0_alice prob1_def mat_of_cols_list_def post_meas1_index_0_alice_state assms(1) a0 a1 select_index_3_subsets by(auto simp add: norm_divide power_divide phi_vec_length algebra_simps) have "(p, q) = ((prob0 3 (alice \) 0) * (prob0 3 (post_meas0 3 (alice \) 0) 1), post_meas0 3 (post_meas0 3 (alice \) 0) 1) \ (p, q) = ((prob0 3 (alice \) 0) * (prob1 3 (post_meas0 3 (alice \) 0) 1), post_meas1 3 (post_meas0 3 (alice \) 0) 1) \ (p, q) = ((prob1 3 (alice \) 0) * (prob0 3 (post_meas1 3 (alice \) 0) 1), post_meas0 3 (post_meas1 3 (alice \) 0) 1) \ (p, q) = ((prob1 3 (alice \) 0) * (prob1 3 (post_meas1 3 (alice \) 0) 1), post_meas1 3 (post_meas1 3 (alice \) 0) 1)" - using assms(3) alice_meas_def List.member_def by(smt list.distinct(1) list.exhaust list.inject member_rec(1) member_rec(2)) + using assms(3) alice_meas_def List.member_def by(smt (verit) list.distinct(1) list.exhaust list.inject member_rec(1) member_rec(2)) then have "q = post_meas0 3 (post_meas0 3 (alice \) 0) 1 \ q = post_meas1 3 (post_meas0 3 (alice \) 0) 1 \ q = post_meas0 3 (post_meas1 3 (alice \) 0) 1 \ q = post_meas1 3 (post_meas1 3 (alice \) 0) 1" by auto moreover have "post_meas0 3 (post_meas0 3 (alice \) 0) 1 = mat_of_cols_list 8 [[\,\,0,0,0,0,0,0]]" proof define v where asm:"v = mat_of_cols_list 8 [[\, \, 0, 0, 0, 0, 0, 0]]" then show "dim_row (post_meas0 3 (post_meas0 3 (alice \) 0) 1) = dim_row v" using mat_of_cols_list_def post_meas0_def ket_vec_def by simp show "dim_col (post_meas0 3 (post_meas0 3 (alice \) 0) 1) = dim_col v" using mat_of_cols_list_def post_meas0_def ket_vec_def asm by simp show "\i j. i j post_meas0 3 (post_meas0 3 (alice \) 0) 1 $$ (i,j) = v $$ (i,j)" proof- fix i j assume "i < dim_row v" and "j < dim_col v" then have "i \ {0..<8} \ j = 0" using asm mat_of_cols_list_def by auto then show "post_meas0 3 (post_meas0 3 (alice \) 0) 1 $$ (i,j) = v $$ (i,j)" using post_meas0_index_0_alice assms(1) a0 a1 apply (auto) using post_meas0_def asm mat_of_cols_list_def ket_vec_def select_index_def by (auto simp add: f0 real_sqrt_divide) qed qed moreover have "post_meas1 3 (post_meas0 3 (alice \) 0) 1 = mat_of_cols_list 8 [[0,0,\,\,0,0,0,0]]" proof define v where asm:"v = mat_of_cols_list 8 [[0,0,\,\,0,0,0,0]]" then show "dim_row (post_meas1 3 (post_meas0 3 (alice \) 0) 1) = dim_row v" using mat_of_cols_list_def post_meas1_def ket_vec_def asm by auto show "dim_col (post_meas1 3 (post_meas0 3 (alice \) 0) 1) = dim_col v" using mat_of_cols_list_def post_meas1_def ket_vec_def asm by auto show "\i j. i j post_meas1 3 (post_meas0 3 (alice \) 0) 1 $$ (i,j) = v $$ (i,j)" proof- fix i j assume "i < dim_row v" and "j < dim_col v" then have "i \ {0..<8} \ j = 0" using asm mat_of_cols_list_def by auto then show "post_meas1 3 (post_meas0 3 (alice \) 0) 1 $$ (i,j) = v $$ (i,j)" using post_meas0_index_0_alice assms(1) a0 a1 apply (auto) using post_meas1_def asm mat_of_cols_list_def ket_vec_def select_index_def by (auto simp add: f1 real_sqrt_divide) qed qed moreover have "post_meas0 3 (post_meas1 3 (alice \) 0) 1 = mat_of_cols_list 8 [[0,0,0,0,\,-\,0,0]]" proof define v where asm:"v = mat_of_cols_list 8 [[0, 0, 0, 0, \, -\, 0, 0]]" then show "dim_row (post_meas0 3 (post_meas1 3 (alice \) 0) 1) = dim_row v" using mat_of_cols_list_def post_meas0_def ket_vec_def by simp show "dim_col (post_meas0 3 (post_meas1 3 (alice \) 0) 1) = dim_col v" using mat_of_cols_list_def post_meas0_def ket_vec_def asm by simp show "\i j. i j post_meas0 3 (post_meas1 3 (alice \) 0) 1 $$ (i,j) = v $$ (i,j)" proof- fix i j assume "i < dim_row v" and "j < dim_col v" then have "i \ {0..<8} \ j = 0" using asm mat_of_cols_list_def by auto then show "post_meas0 3 (post_meas1 3 (alice \) 0) 1 $$ (i,j) = v $$ (i,j)" using post_meas1_index_0_alice assms(1) a0 a1 apply (auto) using post_meas0_def asm mat_of_cols_list_def ket_vec_def select_index_def by (auto simp add: f2 real_sqrt_divide) qed qed moreover have "post_meas1 3 (post_meas1 3 (alice \) 0) 1 = mat_of_cols_list 8 [[0,0,0,0,0,0,-\,\]]" proof define v where asm:"v = mat_of_cols_list 8 [[0,0,0,0,0,0,-\,\]]" then show "dim_row (post_meas1 3 (post_meas1 3 (alice \) 0) 1) = dim_row v" using mat_of_cols_list_def post_meas1_def ket_vec_def by simp show "dim_col (post_meas1 3 (post_meas1 3 (alice \) 0) 1) = dim_col v" using mat_of_cols_list_def post_meas1_def ket_vec_def asm by simp show "\i j. i j post_meas1 3 (post_meas1 3 (alice \) 0) 1 $$ (i,j) = v $$ (i,j)" proof- fix i j assume "i < dim_row v" and "j < dim_col v" then have "i \ {0..<8} \ j = 0" using asm mat_of_cols_list_def by auto then show "post_meas1 3 (post_meas1 3 (alice \) 0) 1 $$ (i,j) = v $$ (i,j)" using post_meas1_index_0_alice assms(1) a0 a1 apply (auto) using post_meas1_def asm mat_of_cols_list_def ket_vec_def select_index_def by (auto simp add: f3 real_sqrt_divide) qed qed ultimately show ?thesis using alice_pos_def a0 a1 by simp qed datatype bit = zero | one definition alice_out:: "complex Matrix.mat => complex Matrix.mat => bit \ bit" where "alice_out \ q \ if q = mat_of_cols_list 8 [[\ $$ (0,0), \ $$ (1,0), 0, 0, 0, 0, 0, 0]] then (zero, zero) else if q = mat_of_cols_list 8 [[0, 0, \ $$ (1,0), \ $$ (0,0), 0, 0, 0, 0]] then (zero, one) else if q = mat_of_cols_list 8 [[0, 0, 0, 0, \ $$ (0,0), - \ $$ (1,0), 0, 0]] then (one, zero) else if q = mat_of_cols_list 8 [[0, 0, 0, 0, 0, 0, - \ $$ (1,0), \ $$ (0,0)]] then (one, one) else undefined" definition bob:: "complex Matrix.mat => bit \ bit \ complex Matrix.mat" where "bob q b \ if (fst b, snd b) = (zero, zero) then q else if (fst b, snd b) = (zero, one) then (Id 2 \ X) * q else if (fst b, snd b) = (one, zero) then (Id 2 \ Z) * q else if (fst b, snd b) = (one, one) then (Id 2 \ Z * X) * q else undefined" lemma alice_out_unique [simp]: assumes "state 1 \" shows "Matrix.mat 8 (Suc 0) (\(i,j). [[0, 0, \ $$ (Suc 0, 0), \ $$ (0, 0), 0, 0, 0, 0]]!j!i) \ Matrix.mat 8 (Suc 0) (\(i,j). [[\ $$ (0, 0), \ $$ (Suc 0, 0), 0, 0, 0, 0, 0, 0]]!j!i) \ Matrix.mat 8 (Suc 0) (\(i,j). [[0, 0, 0, 0, \ $$ (0, 0), -\ $$ (Suc 0, 0), 0, 0]]!j!i) \ Matrix.mat 8 (Suc 0) (\(i,j). [[\ $$ (0, 0), \ $$ (Suc 0, 0), 0, 0, 0, 0, 0, 0]]!j!i) \ Matrix.mat 8 (Suc 0) (\(i,j). [[0, 0, 0, 0, 0, 0, -\ $$ (Suc 0, 0), \ $$ (0, 0)]]!j!i) \ Matrix.mat 8 (Suc 0) (\(i,j). [[\ $$ (0, 0), \ $$ (Suc 0, 0), 0, 0, 0, 0, 0, 0]]!j!i) \ Matrix.mat 8 (Suc 0) (\(i,j). [[0, 0, 0, 0, \ $$ (0, 0), -\ $$ (Suc 0, 0), 0, 0]]!j!i) \ Matrix.mat 8 (Suc 0) (\(i,j). [[0, 0, \ $$ (Suc 0, 0), \ $$ (0, 0), 0, 0, 0, 0]]!j!i) \ Matrix.mat 8 (Suc 0) (\(i,j). [[0, 0, 0, 0, 0, 0, -\ $$ (Suc 0, 0), \ $$ (0, 0)]]!j!i) \ Matrix.mat 8 (Suc 0) (\(i,j). [[0, 0, \ $$ (Suc 0, 0), \ $$ (0, 0), 0, 0, 0, 0]]!j!i) \ Matrix.mat 8 (Suc 0) (\(i,j). [[0, 0, 0, 0, 0, 0, -\ $$ (Suc 0, 0), \ $$ (0, 0)]]!j!i) \ Matrix.mat 8 (Suc 0) (\(i,j). [[0, 0, 0, 0, \ $$ (0, 0), -\ $$ (Suc 0, 0), 0, 0]]!j!i)" proof- have f0:"\ $$ (0,0) \ 0 \ \ $$ (1,0) \ 0" using assms state_def Matrix.col_def cpx_vec_length_def set_2 by(auto simp add: atLeast0LessThan) define v1 v2 v3 v4 where d0:"v1 = Matrix.mat 8 1 (\(i,j). [[\ $$ (0,0),\ $$ (1,0),0,0,0,0,0,0]]!j!i)" and d1:"v2 = Matrix.mat 8 1 (\(i,j). [[0,0,\ $$ (1,0), \ $$ (0,0),0,0,0,0]]!j!i)" and d2:"v3 = Matrix.mat 8 1 (\(i,j). [[0,0,0,0,\ $$ (0,0),-\ $$ (1,0),0,0]]!j!i)" and d3:"v4 = Matrix.mat 8 1 (\(i,j). [[0,0,0,0,0,0,-\ $$ (1,0),\ $$ (0,0)]]!j!i)" then have "(v1 $$ (0,0) \ v2 $$ (0,0) \ v1 $$ (1,0) \ v2 $$ (1,0)) \ (v1 $$ (0,0) \ v3 $$ (0,0) \ v1 $$ (1,0) \ v3 $$ (1,0)) \ (v1 $$ (0,0) \ v4 $$ (0,0) \ v1 $$ (1,0) \ v4 $$ (1,0)) \ (v2 $$ (2,0) \ v3 $$ (2,0) \ v2 $$ (3,0) \ v3 $$ (3,0)) \ (v2 $$ (2,0) \ v4 $$ (2,0) \ v2 $$ (3,0) \ v4 $$ (3,0)) \ (v3 $$ (4,0) \ v4 $$ (4,0) \ v3 $$ (5,0) \ v4 $$ (5,0))" using f0 by auto then have "v1 \ v2 \ v1 \ v3 \ v1 \ v4 \ v2 \ v3 \ v2 \ v4 \ v3 \ v4" by auto thus ?thesis by (auto simp add: d0 d1 d2 d3) qed abbreviation M3:: "complex Matrix.mat" where "M3 \ mat_of_cols_list 8 [[0, 1, 0, 0, 0, 0, 0, 0], [1, 0, 0, 0, 0, 0, 0, 0], [0, 0, 0, 1, 0, 0, 0, 0], [0, 0, 1, 0, 0, 0, 0, 0], [0, 0, 0, 0, 0, 1, 0, 0], [0, 0, 0, 0, 1, 0, 0, 0], [0, 0, 0, 0, 0, 0, 0, 1], [0, 0, 0, 0, 0, 0, 1, 0]]" lemma tensor_prod_of_id_2_x: shows "(Id 2 \ X) = M3" proof have f0:"gate 3 (Id 2 \ X)" using X_is_gate tensor_gate[of "2" "Id 2" "1" "X"] by simp then show "dim_row (Id 2 \ X) = dim_row M3" using gate_def by (simp add: mat_of_cols_list_def) show "dim_col (Id 2 \ X) = dim_col M3" using f0 gate_def by (simp add: mat_of_cols_list_def) show "\i j. i < dim_row M3 \ j < dim_col M3 \ (Id 2 \ X) $$ (i,j) = M3 $$ (i,j)" proof- fix i j assume "i < dim_row M3" and "j < dim_col M3" then have "i \ {0..<8} \ j \ {0..<8}" by (auto simp add: mat_of_cols_list_def) then show "(Id 2 \ X) $$ (i,j) = M3 $$ (i,j)" using Id_def X_def index_tensor_mat[of "Id 2" "4" "4" "X" "2" "2" "i" "j"] gate_def X_is_gate id_is_gate Id_def by (auto simp add: mat_of_cols_list_def X_def) qed qed abbreviation M4:: "complex Matrix.mat" where "M4 \ mat_of_cols_list 8 [[0, -1, 0, 0, 0, 0, 0, 0], [1, 0, 0, 0, 0, 0, 0, 0], [0, 0, 0, -1, 0, 0, 0, 0], [0, 0, 1, 0, 0, 0, 0, 0], [0, 0, 0, 0, 0, -1, 0, 0], [0, 0, 0, 0, 1, 0, 0, 0], [0, 0, 0, 0, 0, 0, 0, -1], [0, 0, 0, 0, 0, 0, 1, 0]]" abbreviation ZX:: "complex Matrix.mat" where "ZX \ mat_of_cols_list 2 [[0, -1], [1, 0]]" lemma l_inv_of_ZX: shows "ZX\<^sup>\ * ZX = 1\<^sub>m 2" proof show "dim_row (ZX\<^sup>\ * ZX) = dim_row (1\<^sub>m 2)" using dagger_def mat_of_cols_list_def by simp show "dim_col (ZX\<^sup>\ * ZX) = dim_col (1\<^sub>m 2)" using dagger_def mat_of_cols_list_def by simp show "\i j. i < dim_row (1\<^sub>m 2) \ j < dim_col (1\<^sub>m 2) \ (ZX\<^sup>\ * ZX) $$ (i, j) = 1\<^sub>m 2 $$ (i, j)" proof- fix i j assume "i < dim_row (1\<^sub>m 2)" and "j < dim_col (1\<^sub>m 2)" then have "i \ {0..<2} \ j \ {0..<2}" by auto then show "(ZX\<^sup>\ * ZX) $$ (i, j) = 1\<^sub>m 2 $$ (i, j)" using mat_of_cols_list_def dagger_def by (auto simp add: set_2) qed qed lemma r_inv_of_ZX: shows "ZX * (ZX\<^sup>\) = 1\<^sub>m 2" proof show "dim_row (ZX * (ZX\<^sup>\)) = dim_row (1\<^sub>m 2)" using dagger_def mat_of_cols_list_def by simp show "dim_col (ZX * (ZX\<^sup>\)) = dim_col (1\<^sub>m 2)" using dagger_def mat_of_cols_list_def by simp show "\i j. i < dim_row (1\<^sub>m 2) \ j < dim_col (1\<^sub>m 2) \ (ZX * (ZX\<^sup>\)) $$ (i, j) = 1\<^sub>m 2 $$ (i, j)" proof- fix i j assume "i < dim_row (1\<^sub>m 2)" and "j < dim_col (1\<^sub>m 2)" then have "i \ {0..<2} \ j \ {0..<2}" by auto then show "(ZX * (ZX\<^sup>\)) $$ (i, j) = 1\<^sub>m 2 $$ (i, j)" using mat_of_cols_list_def dagger_def by (auto simp add: set_2) qed qed lemma ZX_is_gate [simp]: shows "gate 1 ZX" proof show "dim_row ZX = 2 ^ 1" using mat_of_cols_list_def by simp show "square_mat ZX" using mat_of_cols_list_def by simp show "unitary ZX" using unitary_def l_inv_of_ZX r_inv_of_ZX mat_of_cols_list_def by auto qed lemma prod_of_ZX: shows "Z * X = ZX" proof show "dim_row (Z * X) = dim_row ZX" using Z_is_gate mat_of_cols_list_def gate_def by auto show "dim_col (Z * X) = dim_col ZX" using X_is_gate mat_of_cols_list_def gate_def by auto show "\i j. i < dim_row ZX \ j < dim_col ZX \ (Z * X) $$ (i, j) = ZX $$ (i, j)" proof- fix i j assume "i < dim_row ZX" and "j < dim_col ZX" then have "i \ {0..<2} \ j \ {0..<2}" by (auto simp add: mat_of_cols_list_def) then show "(Z * X) $$ (i, j) = ZX $$ (i, j)" by (auto simp add: set_2 Z_def X_def) qed qed lemma tensor_prod_of_id_2_y: shows "(Id 2 \ Z * X) = M4" proof have f0:"gate 3 (Id 2 \ Z * X)" using prod_of_ZX ZX_is_gate tensor_gate[of "2" "Id 2" "1" "Z * X"] by simp then show "dim_row (Id 2 \ Z * X) = dim_row M4" using gate_def by (simp add: mat_of_cols_list_def) show "dim_col (Id 2 \ Z * X) = dim_col M4" using f0 gate_def by (simp add: mat_of_cols_list_def) show "\i j. i < dim_row M4 \ j < dim_col M4 \ (Id 2 \ Z * X) $$ (i,j) = M4 $$ (i,j)" proof- fix i j assume "i < dim_row M4" and "j < dim_col M4" then have "i \ {0..<8} \ j \ {0..<8}" by (auto simp add: mat_of_cols_list_def) then have "(Id 2 \ ZX) $$ (i, j) = M4 $$ (i,j)" using Id_def mat_of_cols_list_def index_tensor_mat[of "Id 2" "4" "4" "ZX" "2" "2" "i" "j"] gate_def ZX_is_gate id_is_gate by (auto simp add: mat_of_cols_list_def) then show "(Id 2 \ Z * X) $$ (i, j) = M4 $$ (i,j)" using prod_of_ZX by simp qed qed abbreviation M5:: "complex Matrix.mat" where "M5 \ mat_of_cols_list 8 [[1, 0, 0, 0, 0, 0, 0, 0], [0, -1, 0, 0, 0, 0, 0, 0], [0, 0, 1, 0, 0, 0, 0, 0], [0, 0, 0, -1, 0, 0, 0, 0], [0, 0, 0, 0, 1, 0, 0, 0], [0, 0, 0, 0, 0, -1, 0, 0], [0, 0, 0, 0, 0, 0, 1, 0], [0, 0, 0, 0, 0, 0, 0, -1]]" lemma tensor_prod_of_id_2_z: shows "(Id 2 \ Z) = M5" proof have f0:"gate 3 (Id 2 \ Z)" using Z_is_gate tensor_gate[of "2" "Id 2" "1" "Z"] by simp then show "dim_row (Id 2 \ Z) = dim_row M5" using gate_def by (simp add: mat_of_cols_list_def) show "dim_col (Id 2 \ Z) = dim_col M5" using f0 gate_def by (simp add: mat_of_cols_list_def) show "\i j. i < dim_row M5 \ j < dim_col M5 \ (Id 2 \ Z) $$ (i,j) = M5 $$ (i,j)" proof- fix i j assume "i < dim_row M5" and "j < dim_col M5" then have "i \ {0..<8} \ j \ {0..<8}" by (auto simp add: mat_of_cols_list_def) then show "(Id 2 \ Z) $$ (i, j) = M5 $$ (i,j)" using Id_def Z_def index_tensor_mat[of "Id 2" "4" "4" "Z" "2" "2" "i" "j"] gate_def Z_is_gate id_is_gate Id_def by (auto simp add: mat_of_cols_list_def Z_def) qed qed lemma teleportation: assumes "state 1 \" and "state 3 q" and "List.member (alice_meas \) (p, q)" shows "\r. state 2 r \ bob q (alice_out \ q) = r \ \" proof- define \ \ where a0:"\ = \ $$ (0,0)" and a1:"\ = \ $$ (1,0)" then have "q = mat_of_cols_list 8 [[\, \, 0, 0, 0, 0, 0, 0]] \ q = mat_of_cols_list 8 [[0, 0, \, \, 0, 0, 0, 0]] \ q = mat_of_cols_list 8 [[0, 0, 0, 0, \, -\, 0, 0]] \ q = mat_of_cols_list 8 [[0, 0, 0, 0, 0, 0, -\, \]]" using assms Alice_case alice_pos_def by simp moreover have "q = mat_of_cols_list 8 [[\,\,0,0,0,0,0,0]] \ bob q (alice_out \ q) = mat_of_cols_list 4 [[1, 0, 0, 0]] \ \" proof assume asm:"q = Tensor.mat_of_cols_list 8 [[\, \, 0, 0, 0, 0, 0, 0]]" show "dim_row (bob q (alice_out \ q)) = dim_row (Tensor.mat_of_cols_list 4 [[1,0,0,0]] \ \)" using mat_of_cols_list_def a0 a1 assms(1) state_def bob_def alice_out_def asm by simp show "dim_col (bob q (alice_out \ q)) = dim_col (Tensor.mat_of_cols_list 4 [[1,0,0,0]] \ \)" using mat_of_cols_list_def a0 a1 assms(1) state_def bob_def alice_out_def asm by simp show "\i j. i < dim_row (Tensor.mat_of_cols_list 4 [[1,0,0,0]] \ \) \ j < dim_col (Tensor.mat_of_cols_list 4 [[1,0,0,0]] \ \) \ bob q (alice_out \ q) $$ (i, j) = (Tensor.mat_of_cols_list 4 [[1,0,0,0]] \ \) $$ (i,j)" proof- fix i j assume "i < dim_row (Tensor.mat_of_cols_list 4 [[1,0,0,0]] \ \)" and "j < dim_col (Tensor.mat_of_cols_list 4 [[1,0,0,0]] \ \)" then have "i \ {0..<8} \ j = 0" using asm mat_of_cols_list_def assms state_def by auto then show "bob q (alice_out \ q) $$ (i,j) = (Tensor.mat_of_cols_list 4 [[1,0,0,0]] \ \) $$ (i,j)" using bob_def alice_out_def asm mat_of_cols_list_def a0 a1 assms state_def by auto qed qed moreover have "q = mat_of_cols_list 8 [[0,0,\,\,0,0,0,0]] \ bob q (alice_out \ q) = mat_of_cols_list 4 [[0,1,0,0]] \ \" proof assume asm:"q = Tensor.mat_of_cols_list 8 [[0,0,\,\,0,0,0,0]]" show "dim_row (bob q (alice_out \ q)) = dim_row (Tensor.mat_of_cols_list 4 [[0,1,0,0]] \ \)" using mat_of_cols_list_def a0 a1 assms(1) state_def bob_def alice_out_def asm tensor_prod_of_id_2_x by simp show "dim_col (bob q (alice_out \ q)) = dim_col (Tensor.mat_of_cols_list 4 [[0,1,0,0]] \ \)" using mat_of_cols_list_def a0 a1 assms(1) state_def bob_def alice_out_def asm by simp show "\i j. i < dim_row (Tensor.mat_of_cols_list 4 [[0,1,0,0]] \ \) \ j < dim_col (Tensor.mat_of_cols_list 4 [[0,1,0,0]] \ \) \ bob q (alice_out \ q) $$ (i,j) = (Tensor.mat_of_cols_list 4 [[0,1,0,0]] \ \) $$ (i,j)" proof- fix i j assume "i < dim_row (Tensor.mat_of_cols_list 4 [[0,1,0,0]] \ \)" and "j < dim_col (Tensor.mat_of_cols_list 4 [[0,1,0,0]] \ \)" then have c1:"i \ {0..<8} \ j = 0" using asm mat_of_cols_list_def assms(1) state_def by auto then have "(M3 * (Matrix.mat 8 1 (\(i,j). [[0,0,\ $$ (1,0),\ $$ (0,0),0,0,0,0]]!j!i))) $$ (i,j) = (Tensor.mat_of_cols_list 4 [[0,1,0,0]] \ \) $$ (i,j)" using state_def assms(1) by(auto simp add: a0 a1 mat_of_cols_list_def times_mat_def scalar_prod_def) then show "bob q (alice_out \ q) $$ (i,j) = (Tensor.mat_of_cols_list 4 [[0,1,0,0]] \ \) $$ (i,j)" using bob_def alice_out_def asm c1 a0 a1 mat_of_cols_list_def tensor_prod_of_id_2_x assms(1) by simp qed qed moreover have "q = mat_of_cols_list 8 [[0,0,0,0,\,-\,0,0]] \ bob q (alice_out \ q) = mat_of_cols_list 4 [[0,0,1,0]] \ \" proof assume asm:"q = Tensor.mat_of_cols_list 8 [[0,0,0,0,\,-\,0,0]]" show "dim_row (bob q (alice_out \ q)) = dim_row (Tensor.mat_of_cols_list 4 [[0,0,1,0]] \ \)" using mat_of_cols_list_def a0 a1 assms(1) state_def bob_def alice_out_def asm tensor_prod_of_id_2_z by simp show "dim_col (bob q (alice_out \ q)) = dim_col (Tensor.mat_of_cols_list 4 [[0,0,1,0]] \ \)" using mat_of_cols_list_def a0 a1 assms(1) state_def bob_def alice_out_def asm by simp show "\i j. i < dim_row (Tensor.mat_of_cols_list 4 [[0,0,1,0]] \ \) \ j < dim_col (Tensor.mat_of_cols_list 4 [[0,0,1,0]] \ \) \ bob q (alice_out \ q) $$ (i,j) = (Tensor.mat_of_cols_list 4 [[0,0,1,0]] \ \) $$ (i,j)" proof- fix i j assume "i < dim_row (Tensor.mat_of_cols_list 4 [[0,0,1,0]] \ \)" and "j < dim_col (Tensor.mat_of_cols_list 4 [[0,0,1,0]] \ \)" then have c1:"i \ {0..<8} \ j = 0" using asm mat_of_cols_list_def assms state_def by auto then have "(M5 * (Matrix.mat 8 (Suc 0) (\(i,j). [[0,0,0,0,\ $$ (0,0),-\ $$ (Suc 0,0),0,0]]!j!i))) $$ (i,j) = (Tensor.mat_of_cols_list 4 [[0,0,1,0]] \ \) $$ (i,j)" using state_def assms(1) by(auto simp add: a0 a1 mat_of_cols_list_def times_mat_def scalar_prod_def) then show "bob q (alice_out \ q) $$ (i,j) = (Tensor.mat_of_cols_list 4 [[0,0,1,0]] \ \) $$ (i,j)" using bob_def alice_out_def asm c1 a0 a1 mat_of_cols_list_def tensor_prod_of_id_2_z assms(1) by simp qed qed moreover have "q = mat_of_cols_list 8 [[0, 0, 0, 0, 0, 0, -\, \]] \ bob q (alice_out \ q) = mat_of_cols_list 4 [[0, 0, 0, 1]] \ \" proof assume asm:"q = Tensor.mat_of_cols_list 8 [[0, 0, 0, 0, 0, 0, -\, \]]" show "dim_row (bob q (alice_out \ q)) = dim_row (Tensor.mat_of_cols_list 4 [[0,0,0,1]] \ \)" using mat_of_cols_list_def a0 a1 assms(1) state_def bob_def alice_out_def asm tensor_prod_of_id_2_y by simp show "dim_col (bob q (alice_out \ q)) = dim_col (Tensor.mat_of_cols_list 4 [[0,0,0,1]] \ \)" using mat_of_cols_list_def a0 a1 assms(1) state_def bob_def alice_out_def asm by simp show "\i j. i < dim_row (Tensor.mat_of_cols_list 4 [[0,0,0,1]] \ \) \ j < dim_col (Tensor.mat_of_cols_list 4 [[0,0,0,1]] \ \) \ bob q (alice_out \ q) $$ (i,j) = (Tensor.mat_of_cols_list 4 [[0,0,0,1]] \ \) $$ (i,j)" proof- fix i j assume "i < dim_row (Tensor.mat_of_cols_list 4 [[0,0,0,1]] \ \)" and "j < dim_col (Tensor.mat_of_cols_list 4 [[0,0,0,1]] \ \)" then have c1:"i \ {0..<8} \ j = 0" using asm mat_of_cols_list_def assms state_def by auto then have "(M4 * (Matrix.mat 8 (Suc 0) (\(i, j). [[0,0,0,0,0,0,-\ $$ (Suc 0,0),\ $$ (0,0)]]!j!i))) $$ (i,j) = (Tensor.mat_of_cols_list 4 [[0,0,0,1]] \ \) $$ (i,j)" using state_def assms(1) by(auto simp add: a0 a1 mat_of_cols_list_def times_mat_def scalar_prod_def) then show "bob q (alice_out \ q) $$ (i,j) = (Tensor.mat_of_cols_list 4 [[0,0,0,1]] \ \) $$ (i,j)" using bob_def alice_out_def asm c1 a0 a1 mat_of_cols_list_def tensor_prod_of_id_2_y assms(1) by simp qed qed moreover have "state 2 (mat_of_cols_list 4 [[1, 0, 0, 0]])" using state_def mat_of_cols_list_def cpx_vec_length_def lessThan_atLeast0 by simp moreover have "state 2 (mat_of_cols_list 4 [[0, 1, 0, 0]])" using state_def mat_of_cols_list_def cpx_vec_length_def lessThan_atLeast0 by simp moreover have "state 2 (mat_of_cols_list 4 [[0, 0, 1, 0]])" using state_def mat_of_cols_list_def cpx_vec_length_def lessThan_atLeast0 by simp moreover have "state 2 (mat_of_cols_list 4 [[0, 0, 0, 1]])" using state_def mat_of_cols_list_def cpx_vec_length_def lessThan_atLeast0 by simp ultimately show ?thesis by auto qed (* Biblio: @inproceedings{Boender2015FormalizationOQ, title={Formalization of Quantum Protocols using Coq}, author={Jaap Boender and Florian Kamm{\"u}ller and Rajagopal Nagarajan}, booktitle={QPL}, year={2015} } *) end \ No newline at end of file diff --git a/thys/Isabelle_Marries_Dirac/Tensor.thy b/thys/Isabelle_Marries_Dirac/Tensor.thy --- a/thys/Isabelle_Marries_Dirac/Tensor.thy +++ b/thys/Isabelle_Marries_Dirac/Tensor.thy @@ -1,471 +1,470 @@ (* Author: Anthony Bordg, University of Cambridge, apdb3@cam.ac.uk *) section \Tensor Products\ theory Tensor imports Complex_Vectors Matrix_Tensor.Matrix_Tensor Jordan_Normal_Form.Matrix begin text \ There is already a formalization of tensor products in the Archive of Formal Proofs, namely Matrix\_Tensor.thy in Tensor Product of Matrices[1] by T.V.H. Prathamesh, but it does not build on top of the formalization of vectors and matrices given in Matrices, Jordan Normal Forms, and Spectral Radius Theory[2] by René Thiemann and Akihisa Yamada. In the present theory our purpose consists in giving such a formalization. Of course, we will reuse Prathamesh's code as much as possible, and in order to achieve that we formalize some lemmas that translate back and forth between vectors (resp. matrices) seen as lists (resp. lists of lists) and vectors (resp. matrices) as formalized in [2]. \ subsection \The Kronecker Product of Complex Vectors\ definition tensor_vec:: "complex Matrix.vec \ complex Matrix.vec \ complex Matrix.vec" (infixl "\" 63) where "tensor_vec u v \ vec_of_list (mult.vec_vec_Tensor (*) (list_of_vec u) (list_of_vec v))" subsection \The Tensor Product of Complex Matrices\ text \To see a matrix in the sense of [2] as a matrix in the sense of [1], we convert it into its list of column vectors.\ definition mat_to_cols_list:: "complex Matrix.mat \ complex list list" where "mat_to_cols_list A = [[A $$ (i,j) . i <- [0..< dim_row A]] . j <- [0..< dim_col A]]" lemma length_mat_to_cols_list [simp]: "length (mat_to_cols_list A) = dim_col A" by (simp add: mat_to_cols_list_def) lemma length_cols_mat_to_cols_list [simp]: assumes "j < dim_col A" shows "length [A $$ (i,j) . i <- [0..< dim_row A]] = dim_row A" using assms by simp lemma length_row_mat_to_cols_list [simp]: assumes "i < dim_row A" shows "length (row (mat_to_cols_list A) i) = dim_col A" using assms by (simp add: row_def) lemma length_col_mat_to_cols_list [simp]: assumes "j < dim_col A" shows "length (col (mat_to_cols_list A) j) = dim_row A" using assms by (simp add: col_def mat_to_cols_list_def) lemma mat_to_cols_list_is_not_Nil [simp]: assumes "dim_col A > 0" shows "mat_to_cols_list A \ []" using assms by (simp add: mat_to_cols_list_def) text \Link between Matrix\_Tensor.row\_length and Matrix.dim\_row\ lemma row_length_mat_to_cols_list [simp]: assumes "dim_col A > 0" shows "mult.row_length (mat_to_cols_list A) = dim_row A" proof - have "mat_to_cols_list A \ []" by (simp add: assms) then have "mult.row_length (mat_to_cols_list A) = length (hd (mat_to_cols_list A))" using mult.row_length_def[of "1" "(*)"] by (simp add: \\xs. Matrix_Tensor.mult 1 (*) \ mult.row_length xs \ if xs = [] then 0 else length (hd xs)\ mult.intro) thus ?thesis by (simp add: assms mat_to_cols_list_def upt_conv_Cons) qed text \@{term mat_to_cols_list} is a matrix in the sense of @{theory Matrix.Matrix_Legacy}.\ lemma mat_to_cols_list_is_mat [simp]: assumes "dim_col A > 0" shows "mat (mult.row_length (mat_to_cols_list A)) (length (mat_to_cols_list A)) (mat_to_cols_list A)" proof - have "Ball (set (mat_to_cols_list A)) (Matrix_Legacy.vec (mult.row_length (mat_to_cols_list A)))" using assms row_length_mat_to_cols_list mat_to_cols_list_def Ball_def set_def vec_def by fastforce thus ?thesis by(auto simp: mat_def) qed definition mat_of_cols_list:: "nat \ complex list list \ complex Matrix.mat" where "mat_of_cols_list nr cs = Matrix.mat nr (length cs) (\ (i,j). cs ! j ! i)" lemma index_mat_of_cols_list [simp]: assumes "i < nr" and "j < length cs" shows "mat_of_cols_list nr cs $$ (i,j) = cs ! j ! i" by (simp add: assms mat_of_cols_list_def) lemma mat_to_cols_list_to_mat [simp]: "mat_of_cols_list (dim_row A) (mat_to_cols_list A) = A" proof show f1:"dim_row (mat_of_cols_list (dim_row A) (mat_to_cols_list A)) = dim_row A" by (simp add: mat_of_cols_list_def) next show f2:"dim_col (mat_of_cols_list (dim_row A) (mat_to_cols_list A)) = dim_col A" by (simp add: Tensor.mat_of_cols_list_def) next show "\i j. i < dim_row A \ j < dim_col A \ (mat_of_cols_list (dim_row A) (mat_to_cols_list A)) $$ (i, j) = A $$ (i, j)" by (simp add: mat_of_cols_list_def mat_to_cols_list_def) qed lemma plus_mult_cpx [simp]: "plus_mult 1 (*) 0 (+) (a_inv cpx_rng)" apply unfold_locales apply (auto intro: cpx_cring_is_field simp: field_simps) proof - show "\x. x + \\<^bsub>cpx_rng\<^esub> x = 0" using group.r_inv[of "cpx_rng"] cpx_cring_is_field field_def domain_def cpx_rng_def by (metis UNIV_I cring.cring_simprules(17) ordered_semiring_record_simps(1) ordered_semiring_record_simps(11) ordered_semiring_record_simps(12)) show "\x. x + \\<^bsub>cpx_rng\<^esub> x = 0" using group.r_inv[of "cpx_rng"] cpx_cring_is_field field_def domain_def cpx_rng_def by (metis UNIV_I cring.cring_simprules(17) ordered_semiring_record_simps(1) ordered_semiring_record_simps(11) ordered_semiring_record_simps(12)) qed lemma list_to_mat_to_cols_list [simp]: fixes l::"complex list list" assumes "mat nr nc l" shows "mat_to_cols_list (mat_of_cols_list nr l) = l" proof - have "length (mat_to_cols_list (mat_of_cols_list nr l)) = length l" by (simp add: mat_of_cols_list_def) moreover have f1:"\jj. j mat_to_cols_list (mat_of_cols_list nr l) ! j = l ! j" proof fix j assume a:"j < length l" then have f2:"length (mat_to_cols_list (mat_of_cols_list nr l) ! j) = length (l ! j)" by (metis col_def mat_def vec_def mat_of_cols_list_def assms dim_col_mat(1) dim_row_mat(1) length_col_mat_to_cols_list nth_mem) then have "\ii complex Matrix.mat" (infixl "\" 63) where "tensor_mat A B \ mat_of_cols_list (dim_row A * dim_row B) (mult.Tensor (*) (mat_to_cols_list A) (mat_to_cols_list B))" lemma dim_row_tensor_mat [simp]: "dim_row (A \ B) = dim_row A * dim_row B" by (simp add: mat_of_cols_list_def tensor_mat_def) lemma dim_col_tensor_mat [simp]: "dim_col (A \ B) = dim_col A * dim_col B" using tensor_mat_def mat_of_cols_list_def mult.length_Tensor[of "1" "(*)"] by(simp add: \\M2 M1. Matrix_Tensor.mult 1 (*) \ length (mult.Tensor (*) M1 M2) = length M1 * length M2\ mult.intro) lemma index_tensor_mat [simp]: assumes a1:"dim_row A = rA" and a2:"dim_col A = cA" and a3:"dim_row B = rB" and a4:"dim_col B = cB" and a5:"i < rA * rB" and a6:"j < cA * cB" and a7:"cA > 0" and a8:"cB > 0" shows "(A \ B) $$ (i,j) = A $$ (i div rB, j div cB) * B $$ (i mod rB, j mod cB)" proof - have "(A \ B) $$ (i,j) = (mult.Tensor (*) (mat_to_cols_list A) (mat_to_cols_list B)) ! j ! i" using assms tensor_mat_def mat_of_cols_list_def dim_col_tensor_mat by simp moreover have f:"i < mult.row_length (mat_to_cols_list A) * mult.row_length (mat_to_cols_list B)" by (simp add: a1 a2 a3 a4 a5 a7 a8) moreover have "j < length (mat_to_cols_list A) * length (mat_to_cols_list B)" by (simp add: a2 a4 a6) moreover have "mat (mult.row_length (mat_to_cols_list A)) (length (mat_to_cols_list A)) (mat_to_cols_list A)" using a2 a7 mat_to_cols_list_is_mat by blast moreover have "mat (mult.row_length (mat_to_cols_list B)) (length (mat_to_cols_list B)) (mat_to_cols_list B)" using a4 a8 mat_to_cols_list_is_mat by blast ultimately have "(A \ B) $$ (i,j) = (mat_to_cols_list A) ! (j div length (mat_to_cols_list B)) ! (i div mult.row_length (mat_to_cols_list B)) * (mat_to_cols_list B) ! (j mod length (mat_to_cols_list B)) ! (i mod mult.row_length (mat_to_cols_list B))" using mult.matrix_Tensor_elements[of "1" "(*)"] by(simp add: \\M2 M1. mult 1 (*) \ \i j. (i j mat (mult.row_length M1) (length M1) M1 \ mat (mult.row_length M2) (length M2) M2 \ mult.Tensor (*) M1 M2 ! j ! i = M1 ! (j div length M2) ! (i div mult.row_length M2) * M2 ! (j mod length M2) ! (i mod mult.row_length M2)\ mult.intro) thus ?thesis using mat_to_cols_list_def by (metis a2 a3 a4 a6 f index_mat_of_cols_list length_mat_to_cols_list less_mult_imp_div_less less_nat_zero_code mat_to_cols_list_to_mat mult_0_right neq0_conv row_length_mat_to_cols_list unique_euclidean_semiring_numeral_class.pos_mod_bound) qed text \To go from @{term Matrix.row} to @{term Matrix_Legacy.row}\ lemma Matrix_row_is_Legacy_row: assumes "i < dim_row A" shows "Matrix.row A i = vec_of_list (row (mat_to_cols_list A) i)" proof show "dim_vec (Matrix.row A i) = dim_vec (vec_of_list (row (mat_to_cols_list A) i))" using length_mat_to_cols_list Matrix.dim_vec_of_list by (metis row_def index_row(2) length_map) next show "\j. j Matrix.row A i $ j = vec_of_list (row (mat_to_cols_list A) i) $ j" using Matrix.row_def vec_of_list_def mat_to_cols_list_def - by(smt row_def assms dim_vec_of_list index_mat_of_cols_list index_row(1) + by (smt (verit) row_def assms dim_vec_of_list index_mat_of_cols_list index_row(1) length_mat_to_cols_list length_row_mat_to_cols_list mat_to_cols_list_to_mat nth_map vec_of_list_index) qed text \To go from @{term Matrix_Legacy.row} to @{term Matrix.row}\ lemma Legacy_row_is_Matrix_row: assumes "i < mult.row_length A" shows "row A i = list_of_vec (Matrix.row (mat_of_cols_list (mult.row_length A) A) i)" proof (rule nth_equalityI) show "length (row A i) = length (list_of_vec (Matrix.row (mat_of_cols_list (mult.row_length A) A) i))" using row_def length_list_of_vec by(metis mat_of_cols_list_def dim_col_mat(1) index_row(2) length_map) next fix j:: nat assume "j < length (row A i)" then show "row A i ! j = list_of_vec (Matrix.row (mat_of_cols_list (mult.row_length A) A) i) ! j" using assms index_mat_of_cols_list by(metis row_def mat_of_cols_list_def dim_col_mat(1) dim_row_mat(1) index_row(1) length_map list_of_vec_index nth_map) qed text \To go from @{term Matrix.col} to @{term Matrix_Legacy.col}\ lemma Matrix_col_is_Legacy_col: assumes "j < dim_col A" shows "Matrix.col A j = vec_of_list (col (mat_to_cols_list A) j)" proof show "dim_vec (Matrix.col A j) = dim_vec (vec_of_list (col (mat_to_cols_list A) j))" by (simp add: col_def assms mat_to_cols_list_def) next show "\i. i < dim_vec (vec_of_list (col (mat_to_cols_list A) j)) \ Matrix.col A j $ i = vec_of_list (col (mat_to_cols_list A) j) $ i" using mat_to_cols_list_def by (metis col_def assms col_mat_of_cols_list length_col_mat_to_cols_list length_mat_to_cols_list mat_to_cols_list_to_mat) qed text \To go from @{term Matrix_Legacy.col} to @{term Matrix.col}\ lemma Legacy_col_is_Matrix_col: assumes a1:"j < length A" and a2:"length (A ! j) = mult.row_length A" shows "col A j = list_of_vec (Matrix.col (mat_of_cols_list (mult.row_length A) A) j)" proof (rule nth_equalityI) have "length (list_of_vec (Matrix.col (mat_of_cols_list (mult.row_length A) A) j)) = dim_vec (Matrix.col (mat_of_cols_list (mult.row_length A) A) j)" using length_list_of_vec by blast also have "\ = dim_row (mat_of_cols_list (mult.row_length A) A)" using Matrix.col_def by simp also have f1:"\ = mult.row_length A" by (simp add: mat_of_cols_list_def) finally show f2:"length (col A j) = length (list_of_vec (Matrix.col (mat_of_cols_list (mult.row_length A) A) j))" using a2 by (simp add: col_def) next fix i:: nat assume "iLink between @{term plus_mult.scalar_product} and @{term Matrix.scalar_prod}\ lemma scalar_prod_is_Matrix_scalar_prod [simp]: fixes u::"complex list" and v::"complex list" assumes "length u = length v" shows "plus_mult.scalar_product (*) 0 (+) u v = (vec_of_list u) \ (vec_of_list v)" proof - have f:"(vec_of_list u) \ (vec_of_list v) = (\i=0.. = sum_list (map (\(x,y). x * y) (zip u v))" by (simp add: scalar_prod) moreover have "\ii(x,y). x * y) (zip u v)) ! i = u ! i * v ! i" by (simp add: assms) ultimately have "plus_mult.scalar_product (*) 0 (+) u v = (\i=0..Link between @{term times} and @{term plus_mult.matrix_mult}\ lemma matrix_mult_to_times_mat: assumes "dim_col A > 0" and "dim_col B > 0" and "dim_col (A::complex Matrix.mat) = dim_row B" shows "A * B = mat_of_cols_list (dim_row A) (plus_mult.matrix_mult (*) 0 (+) (mat_to_cols_list A) (mat_to_cols_list B))" proof define M where "M = mat_of_cols_list (dim_row A) (plus_mult.matrix_mult (*) 0 (+) (mat_to_cols_list A) (mat_to_cols_list B))" then show f1:"dim_row (A * B) = dim_row M" by (simp add: mat_of_cols_list_def times_mat_def) have "length (plus_mult.matrix_mult (*) 0 (+) (mat_to_cols_list A) (mat_to_cols_list B)) = dim_col B" by (simp add: mat_multI_def) then show f2:"dim_col (A * B) = dim_col M" by (simp add: M_def times_mat_def mat_of_cols_list_def) show "\i j. i < dim_row M \ j < dim_col M \ (A * B) $$ (i, j) = M $$ (i, j)" proof - fix i j assume a1:"i < dim_row M" and a2:"j < dim_col M" then have "(A * B) $$ (i,j) = Matrix.row A i \ Matrix.col B j" using f1 f2 by simp also have "\ = vec_of_list (row (mat_to_cols_list A) i) \ vec_of_list (col (mat_to_cols_list B) j)" using f1 f2 a1 a2 by (simp add: Matrix_row_is_Legacy_row Matrix_col_is_Legacy_col) also have "\ = plus_mult.scalar_product (*) 0 (+) (row (mat_to_cols_list A) i) (col (mat_to_cols_list B) j)" using a1 a2 assms(3) f1 f2 by simp also have "M $$ (i,j) = plus_mult.scalar_product (*) 0 (+) (row (mat_to_cols_list A) i) (col (mat_to_cols_list B) j)" proof- have "M $$ (i,j) = (plus_mult.matrix_mult (*) 0 (+) (mat_to_cols_list A) (mat_to_cols_list B)) ! j ! i" using M_def f1 f2 \length (mat_mult (mult.row_length (mat_to_cols_list A)) (mat_to_cols_list A) (mat_to_cols_list B)) = dim_col B\ a1 a2 by simp moreover have "mat (mult.row_length (mat_to_cols_list A)) (dim_col A) (mat_to_cols_list A)" using mat_to_cols_list_is_mat assms(1) by simp moreover have "mat (dim_col A) (dim_col B) (mat_to_cols_list B)" using assms(2) assms(3) mat_to_cols_list_is_mat by simp ultimately show ?thesis using assms(1) a1 a2 row_length_mat_to_cols_list plus_mult.matrix_index[of 1 "(*)" 0 "(+)"] plus_mult_cpx - by (smt f1 f2 index_mult_mat(2) index_mult_mat(3)) + by (smt (verit) f1 f2 index_mult_mat(2) index_mult_mat(3)) qed finally show "(A * B) $$ (i, j) = M $$ (i, j)" by simp qed qed lemma mat_to_cols_list_times_mat [simp]: assumes "dim_col A = dim_row B" and "dim_col A > 0" shows "mat_to_cols_list (A * B) = plus_mult.matrix_mult (*) 0 (+) (mat_to_cols_list A) (mat_to_cols_list B)" proof (rule nth_equalityI) define M where "M = plus_mult.matrix_mult (*) 0 (+) (mat_to_cols_list A) (mat_to_cols_list B)" then show f0:"length (mat_to_cols_list (A * B)) = length M" by (simp add: mat_multI_def) moreover have f1:"\j. j mat_to_cols_list (A * B) ! j = M ! j" proof fix j:: nat assume a0:"j < length (mat_to_cols_list (A * B))" then have "length (mat_to_cols_list (A * B) ! j) = dim_row A" by (simp add: mat_to_cols_list_def) then also have f2:"length (M ! j) = dim_row A" using a0 M_def mat_multI_def[of 0 "(+)" "(*)" "dim_row A" "mat_to_cols_list A" "mat_to_cols_list B"] row_length_mat_to_cols_list assms(2) by (metis assms(1) f0 length_greater_0_conv length_map length_mat_to_cols_list list_to_mat_to_cols_list mat_mult mat_to_cols_list_is_mat matrix_mult_to_times_mat) ultimately have "length (mat_to_cols_list (A * B) ! j) = length (M ! j)" by simp moreover have "\i. i mat_to_cols_list (A * B) ! j ! i = M ! j ! i" proof fix i assume a1:"i < dim_row A" have "mat (mult.row_length (mat_to_cols_list A)) (dim_col A) (mat_to_cols_list A)" using mat_to_cols_list_is_mat assms(2) by simp moreover have "mat (dim_col A) (dim_col B) (mat_to_cols_list B)" using mat_to_cols_list_is_mat assms(1) a0 by simp ultimately have "M ! j ! i = plus_mult.scalar_product (*) 0 (+) (row (mat_to_cols_list A) i) (col (mat_to_cols_list B) j)" using plus_mult.matrix_index a0 a1 row_length_mat_to_cols_list assms(2) plus_mult_cpx M_def by (metis index_mult_mat(3) length_mat_to_cols_list) also have "\ = vec_of_list (row (mat_to_cols_list A) i) \ vec_of_list (col (mat_to_cols_list B) j)" using a0 a1 assms(1) by simp finally show "mat_to_cols_list (A * B) ! j ! i = M ! j ! i" using mat_to_cols_list_def index_mult_mat(1) a0 a1 by(simp add: Matrix_row_is_Legacy_row Matrix_col_is_Legacy_col) qed ultimately show "mat_to_cols_list (A * B) ! j = M ! j" by(simp add: nth_equalityI f2) qed fix i:: nat assume "i < length (mat_to_cols_list (A * B))" thus "mat_to_cols_list (A * B) ! i = M ! i" by (simp add: f1) qed text \ Finally, we prove that the tensor product of complex matrices is distributive over the multiplication of complex matrices. \ lemma mult_distr_tensor: assumes a1:"dim_col A = dim_row B" and a2:"dim_col C = dim_row D" and a3:"dim_col A > 0" and a4:"dim_col B > 0" and a5:"dim_col C > 0" and a6:"dim_col D > 0" shows "(A * B) \ (C * D) = (A \ C) * (B \ D)" proof - define A' B' C' D' M N where "A' = mat_to_cols_list A" and "B' = mat_to_cols_list B" and "C' = mat_to_cols_list C" and "D' = mat_to_cols_list D" and "M = mat_of_cols_list (dim_row A * dim_row C) (mult.Tensor (*) (mat_to_cols_list A) (mat_to_cols_list C))" and "N = mat_of_cols_list (dim_row B * dim_row D) (mult.Tensor (*) (mat_to_cols_list B) (mat_to_cols_list D))" then have "(A \ C) * (B \ D) = M * N" by (simp add: tensor_mat_def) also have "\ = mat_of_cols_list (dim_row A * dim_row C) (plus_mult.matrix_mult (*) 0 (+) (mat_to_cols_list M) (mat_to_cols_list N))" using assms M_def N_def dim_col_tensor_mat dim_row_tensor_mat tensor_mat_def by(simp add: matrix_mult_to_times_mat) also have f1:"\ = mat_of_cols_list (dim_row A * dim_row C) (plus_mult.matrix_mult (*) 0 (+) (mult.Tensor (*) A' C') (mult.Tensor (*) B' D'))" proof - define M' N' where "M' = mult.Tensor (*) (mat_to_cols_list A) (mat_to_cols_list C)" and "N' = mult.Tensor (*) (mat_to_cols_list B) (mat_to_cols_list D)" then have "mat (mult.row_length M') (length M') M'" using M'_def mult.effective_well_defined_Tensor[of 1 "(*)"] mat_to_cols_list_is_mat a3 a5 - by (smt mult.length_Tensor mult.row_length_mat plus_mult_cpx plus_mult_def) + by (smt (verit) mult.length_Tensor mult.row_length_mat plus_mult_cpx plus_mult_def) moreover have "mat (mult.row_length N') (length N') N'" using N'_def mult.effective_well_defined_Tensor[of 1 "(*)"] mat_to_cols_list_is_mat a4 a6 - by (smt mult.length_Tensor mult.row_length_mat plus_mult_cpx plus_mult_def) + by (smt (verit) mult.length_Tensor mult.row_length_mat plus_mult_cpx plus_mult_def) ultimately show ?thesis using list_to_mat_to_cols_list M_def N_def mult.row_length_mat row_length_mat_to_cols_list assms(3) a4 a5 a6 A'_def B'_def C'_def D'_def by(metis M'_def N'_def plus_mult_cpx plus_mult_def) qed also have "\ = mat_of_cols_list (dim_row A * dim_row C) (mult.Tensor (*) (plus_mult.matrix_mult (*) 0 (+) A' B') (plus_mult.matrix_mult (*) 0 (+) C' D'))" proof - have f2:"mat (mult.row_length A') (length A') A'" using A'_def a3 mat_to_cols_list_is_mat by simp moreover have "mat (mult.row_length B') (length B') B'" using B'_def a4 mat_to_cols_list_is_mat by simp moreover have "mat (mult.row_length C') (length C') C'" using C'_def a5 mat_to_cols_list_is_mat by simp moreover have "mat (mult.row_length D') (length D') D'" using D'_def a6 mat_to_cols_list_is_mat by simp moreover have "length A' = mult.row_length B'" using A'_def B'_def a1 a4 by simp moreover have "length C' = mult.row_length D'" using C'_def D'_def a2 a6 by simp moreover have "A' \ [] \ B' \ [] \ C' \ [] \ D' \ []" using A'_def B'_def C'_def D'_def a3 a4 a5 a6 by simp ultimately have "plus_mult.matrix_match A' B' C' D'" using plus_mult.matrix_match_def[of 1 "(*)" 0 "(+)" "a_inv cpx_rng"] by simp thus ?thesis using f1 plus_mult.distributivity plus_mult_cpx by fastforce qed also have "\ = mat_of_cols_list (dim_row A * dim_row C) (mult.Tensor (*) (mat_to_cols_list (A * B)) (mat_to_cols_list (C * D)))" using A'_def B'_def C'_def D'_def a1 a2 a3 a5 by simp finally show ?thesis by(simp add: tensor_mat_def) qed lemma tensor_mat_is_assoc: fixes A B C:: "complex Matrix.mat" shows "A \ (B \ C) = (A \ B) \ C" proof- define M where d:"M = mat_of_cols_list (dim_row B * dim_row C) (mult.Tensor (*) (mat_to_cols_list B) (mat_to_cols_list C))" - then have "B \ C = M" + then have *: "B \ C = M" using tensor_mat_def by simp - moreover have "A \ (B \ C) = mat_of_cols_list (dim_row A * (dim_row B * dim_row C)) + then have **: "A \ (B \ C) = mat_of_cols_list (dim_row A * (dim_row B * dim_row C)) (mult.Tensor (*) (mat_to_cols_list A) (mat_to_cols_list M))" using tensor_mat_def d dim_row_tensor_mat by simp - moreover have "mat_to_cols_list M = mult.Tensor (*) (mat_to_cols_list B) (mat_to_cols_list C)" - using d list_to_mat_to_cols_list - by (smt calculation(1) dim_col_tensor_mat length_greater_0_conv length_mat_to_cols_list mat_to_cols_list_is_mat -mult.Tensor.simps(1) mult.Tensor_null mult.well_defined_Tensor nat_0_less_mult_iff plus_mult_cpx plus_mult_def row_length_mat_to_cols_list) - ultimately have "A \ (B \ C) = mat_of_cols_list (dim_row A * (dim_row B * dim_row C)) + then have ***: "mat_to_cols_list M = mult.Tensor (*) (mat_to_cols_list B) (mat_to_cols_list C)" + unfolding d using list_to_mat_to_cols_list + by (smt (verit) Tensor.mat_of_cols_list_def bot_nat_0.not_eq_extremum * d dim_col_mat(1) dim_col_tensor_mat length_greater_0_conv length_mat_to_cols_list mat_to_cols_list_is_mat mult.effective_well_defined_Tensor mult_is_0 plus_mult_cpx plus_mult_def row_length_mat_to_cols_list) + with ** have "A \ (B \ C) = mat_of_cols_list (dim_row A * (dim_row B * dim_row C)) (mult.Tensor (*) (mat_to_cols_list A) (mult.Tensor (*) (mat_to_cols_list B) (mat_to_cols_list C)))" by simp moreover have "\ = mat_of_cols_list ((dim_row A * dim_row B) * dim_row C) (mult.Tensor (*) (mult.Tensor (*) (mat_to_cols_list A) (mat_to_cols_list B)) (mat_to_cols_list C))" using Matrix_Tensor.mult.associativity - by (smt ab_semigroup_mult_class.mult_ac(1) length_greater_0_conv length_mat_to_cols_list + by (smt (verit) ab_semigroup_mult_class.mult_ac(1) length_greater_0_conv length_mat_to_cols_list mat_to_cols_list_is_mat mult.Tensor.simps(1) mult.Tensor_null plus_mult_cpx plus_mult_def) ultimately show ?thesis using tensor_mat_def - by (smt Tensor.mat_of_cols_list_def dim_col_mat(1) dim_col_tensor_mat dim_row_tensor_mat length_0_conv + by (smt (verit) Tensor.mat_of_cols_list_def dim_col_mat(1) dim_col_tensor_mat dim_row_tensor_mat length_0_conv list_to_mat_to_cols_list mat_to_cols_list_is_mat mult.well_defined_Tensor mult_is_0 neq0_conv plus_mult_cpx plus_mult_def row_length_mat_to_cols_list) qed end \ No newline at end of file