diff --git a/thys/Complex_Geometry/Chordal_Metric.thy b/thys/Complex_Geometry/Chordal_Metric.thy --- a/thys/Complex_Geometry/Chordal_Metric.thy +++ b/thys/Complex_Geometry/Chordal_Metric.thy @@ -1,1725 +1,1713 @@ (* -------------------------------------------------------------------------- *) subsection \Chordal Metric\ (* -------------------------------------------------------------------------- *) text \Riemann sphere can be made a metric space. We are going to introduce distance on Riemann sphere and to prove that it is a metric space. The distance between two points on the sphere is defined as the length of the chord that connects them. This metric can be used in formalization of elliptic geometry.\ theory Chordal_Metric imports Homogeneous_Coordinates Riemann_Sphere Oriented_Circlines "HOL-Analysis.Inner_Product" "HOL-Analysis.Euclidean_Space" begin (* -------------------------------------------------------------------------- *) subsubsection \Inner product and norm\ (* -------------------------------------------------------------------------- *) definition inprod_cvec :: "complex_vec \ complex_vec \ complex" where [simp]: "inprod_cvec z w = (let (z1, z2) = z; (w1, w2) = w in vec_cnj (z1, z2) *\<^sub>v\<^sub>v (w1, w2))" syntax "_inprod_cvec" :: "complex_vec \ complex_vec \ complex" ("\_,_\") translations "\z,w\" == "CONST inprod_cvec z w" lemma real_inprod_cvec [simp]: shows "is_real \z,z\" by (cases z, simp add: vec_cnj_def) lemma inprod_cvec_ge_zero [simp]: shows "Re \z,z\ \ 0" by (cases z, simp add: vec_cnj_def) lemma inprod_cvec_bilinear1 [simp]: assumes "z' = k *\<^sub>s\<^sub>v z" shows "\z',w\ = cnj k * \z,w\" using assms by (cases z, cases z', cases w) (simp add: vec_cnj_def field_simps) lemma inprod_cvec_bilinear2 [simp]: assumes "z' = k *\<^sub>s\<^sub>v z" shows "\w, z'\ = k * \w, z\" using assms by (cases z, cases z', cases w) (simp add: vec_cnj_def field_simps) lemma inprod_cvec_g_zero [simp]: assumes "z \ vec_zero" shows "Re \z, z\ > 0" proof- have "\ a b. a \ 0 \ b \ 0 \ 0 < (Re a * Re a + Im a * Im a) + (Re b * Re b + Im b * Im b)" by (smt complex_eq_0 not_sum_squares_lt_zero power2_eq_square) thus ?thesis using assms by (cases z, simp add: vec_cnj_def) qed definition norm_cvec :: "complex_vec \ real" where [simp]: "norm_cvec z = sqrt (Re \z,z\)" syntax "_norm_cvec" :: "complex_vec \ complex" ("\_\") translations "\z\" == "CONST norm_cvec z" lemma norm_cvec_square: shows "\z\\<^sup>2 = Re (\z,z\)" by (simp del: inprod_cvec_def) lemma norm_cvec_gt_0: assumes "z \ vec_zero" shows "\z\ > 0" using assms by (simp del: inprod_cvec_def) lemma norm_cvec_scale: assumes "z' = k *\<^sub>s\<^sub>v z" shows "\z'\\<^sup>2 = Re (cnj k * k) * \z\\<^sup>2" unfolding norm_cvec_square using inprod_cvec_bilinear1[OF assms, of z'] using inprod_cvec_bilinear2[OF assms, of z] by (simp del: inprod_cvec_def add: field_simps) lift_definition inprod_hcoords :: "complex_homo_coords \ complex_homo_coords \ complex" is inprod_cvec done lift_definition norm_hcoords :: "complex_homo_coords \ real" is norm_cvec done (* -------------------------------------------------------------------------- *) subsubsection \Distance in $\mathbb{C}P^1$ - defined by Fubini-Study metric.\ (* -------------------------------------------------------------------------- *) text \Formula for the chordal distance between the two points can be given directly based on the homogenous coordinates of their stereographic projections in the plane. This is called the Fubini-Study metric.\ definition dist_fs_cvec :: "complex_vec \ complex_vec \ real" where [simp]: "dist_fs_cvec z1 z2 = (let (z1x, z1y) = z1; (z2x, z2y) = z2; num = (z1x*z2y - z2x*z1y) * (cnj z1x*cnj z2y - cnj z2x*cnj z1y); den = (z1x*cnj z1x + z1y*cnj z1y) * (z2x*cnj z2x + z2y*cnj z2y) in 2*sqrt(Re num / Re den))" lemma dist_fs_cvec_iff: assumes "z \ vec_zero" and "w \ vec_zero" shows "dist_fs_cvec z w = 2*sqrt(1 - (cmod \z,w\)\<^sup>2 / (\z\\<^sup>2 * \w\\<^sup>2))" proof- obtain z1 z2 w1 w2 where *: "z = (z1, z2)" "w = (w1, w2)" by (cases "z", cases "w") auto have 1: "2*sqrt(1 - (cmod \z,w\)\<^sup>2 / (\z\\<^sup>2 * \w\\<^sup>2)) = 2*sqrt((\z\\<^sup>2 * \w\\<^sup>2 - (cmod \z,w\)\<^sup>2) / (\z\\<^sup>2 * \w\\<^sup>2))" using norm_cvec_gt_0[of z] norm_cvec_gt_0[of w] assms by (simp add: field_simps) have 2: "\z\\<^sup>2 * \w\\<^sup>2 = Re ((z1*cnj z1 + z2*cnj z2) * (w1*cnj w1 + w2*cnj w2))" using assms * by (simp add: vec_cnj_def) have 3: "\z\\<^sup>2 * \w\\<^sup>2 - (cmod \z,w\)\<^sup>2 = Re ((z1*w2 - w1*z2) * (cnj z1*cnj w2 - cnj w1*cnj z2))" apply (subst cmod_square, (subst norm_cvec_square)+) using * by (simp add: vec_cnj_def field_simps) thus ?thesis using 1 2 3 using * unfolding dist_fs_cvec_def Let_def by simp qed lift_definition dist_fs_hcoords :: "complex_homo_coords \ complex_homo_coords \ real" is dist_fs_cvec done lift_definition dist_fs :: "complex_homo \ complex_homo \ real" is dist_fs_hcoords proof transfer fix z1 z2 z1' z2' :: complex_vec obtain z1x z1y z2x z2y z1'x z1'y z2'x z2'y where zz: "z1 = (z1x, z1y)" "z2 = (z2x, z2y)" "z1' = (z1'x, z1'y)" "z2' = (z2'x, z2'y)" by (cases "z1", cases "z2", cases "z1'", cases "z2'") blast assume 1: "z1 \ vec_zero" "z2 \ vec_zero" "z1' \ vec_zero" "z2' \ vec_zero" "z1 \\<^sub>v z1'" "z2 \\<^sub>v z2'" then obtain k1 k2 where *: "k1 \ 0" "z1' = k1 *\<^sub>s\<^sub>v z1" and **: "k2 \ 0" "z2' = k2 *\<^sub>s\<^sub>v z2" by auto have "(cmod \z1,z2\)\<^sup>2 / (\z1\\<^sup>2 * \z2\\<^sup>2) = (cmod \z1',z2'\)\<^sup>2 / (\z1'\\<^sup>2 * \z2'\\<^sup>2)" using \k1 \ 0\ \k2 \ 0\ using cmod_square[symmetric, of k1] cmod_square[symmetric, of k2] apply (subst norm_cvec_scale[OF *(2)]) apply (subst norm_cvec_scale[OF **(2)]) apply (subst inprod_cvec_bilinear1[OF *(2)]) apply (subst inprod_cvec_bilinear2[OF **(2)]) - by (simp add: power2_eq_square) + by (simp add: power2_eq_square norm_mult) thus "dist_fs_cvec z1 z2 = dist_fs_cvec z1' z2'" using 1 dist_fs_cvec_iff by simp qed lemma dist_fs_finite: shows "dist_fs (of_complex z1) (of_complex z2) = 2 * cmod(z1 - z2) / (sqrt (1+(cmod z1)\<^sup>2) * sqrt (1+(cmod z2)\<^sup>2))" apply transfer apply transfer apply (subst cmod_square)+ apply (simp add: real_sqrt_divide cmod_def power2_eq_square) apply (subst real_sqrt_mult[symmetric]) apply (simp add: field_simps) done lemma dist_fs_infinite1: shows "dist_fs (of_complex z1) \\<^sub>h = 2 / sqrt (1+(cmod z1)\<^sup>2)" by (transfer, transfer) (subst cmod_square, simp add: real_sqrt_divide) lemma dist_fs_infinite2: shows "dist_fs \\<^sub>h (of_complex z1) = 2 / sqrt (1+(cmod z1)\<^sup>2)" by (transfer, transfer) (subst cmod_square, simp add: real_sqrt_divide) lemma dist_fs_cvec_zero: assumes "z \ vec_zero" and "w \ vec_zero" shows "dist_fs_cvec z w = 0 \ (cmod \z,w\)\<^sup>2 = (\z\\<^sup>2 * \w\\<^sup>2)" using assms norm_cvec_gt_0[of z] norm_cvec_gt_0[of w] by (subst dist_fs_cvec_iff) auto lemma dist_fs_zero1 [simp]: shows "dist_fs z z = 0" by (transfer, transfer) (subst dist_fs_cvec_zero, simp, (subst norm_cvec_square)+, subst cmod_square, simp del: inprod_cvec_def) lemma dist_fs_zero2 [simp]: assumes "dist_fs z1 z2 = 0" shows "z1 = z2" using assms proof (transfer, transfer) fix z w :: complex_vec obtain z1 z2 w1 w2 where *: "z = (z1, z2)" "w = (w1, w2)" by (cases "z", cases "w", auto) let ?x = "(z1*w2 - w1*z2) * (cnj z1*cnj w2 - cnj w1*cnj z2)" assume "z \ vec_zero" "w \ vec_zero" "dist_fs_cvec z w = 0" hence "(cmod \z,w\)\<^sup>2 = \z\\<^sup>2 * \w\\<^sup>2" by (subst (asm) dist_fs_cvec_zero, simp_all) hence "Re ?x = 0" using * by (subst (asm) cmod_square) ((subst (asm) norm_cvec_square)+, simp add: vec_cnj_def field_simps) hence "?x = 0" using complex_mult_cnj_cmod[of "z1*w2 - w1*z2"] zero_complex.simps by (subst complex_eq_if_Re_eq[of ?x 0]) (simp add: power2_eq_square, simp, linarith) moreover have "z1 * w2 - w1 * z2 = 0 \ cnj z1 * cnj w2 - cnj w1 * cnj z2 = 0" by (metis complex_cnj_diff complex_cnj_mult complex_cnj_zero_iff) ultimately show "z \\<^sub>v w" using * \z \ vec_zero\ \w \ vec_zero\ using complex_cvec_eq_mix[of z1 z2 w1 w2] by auto qed lemma dist_fs_sym: shows "dist_fs z1 z2 = dist_fs z2 z1" by (transfer, transfer) (simp add: split_def field_simps) (* -------------------------------------------------------------------------- *) subsubsection \Triangle inequality for Fubini-Study metric\ (* -------------------------------------------------------------------------- *) lemma dist_fs_triangle_finite: shows "cmod(a - b) / (sqrt (1+(cmod a)\<^sup>2) * sqrt (1+(cmod b)\<^sup>2)) \ cmod (a - c) / (sqrt (1+(cmod a)\<^sup>2) * sqrt (1+(cmod c)\<^sup>2)) + cmod (c - b) / (sqrt (1+(cmod b)\<^sup>2) * sqrt (1+(cmod c)\<^sup>2))" proof- let ?cc = "1+(cmod c)\<^sup>2" and ?bb = "1+(cmod b)\<^sup>2" and ?aa = "1+(cmod a)\<^sup>2" have "sqrt ?cc > 0" "sqrt ?aa > 0" "sqrt ?bb > 0" by (smt real_sqrt_gt_zero zero_compare_simps(12))+ have "(a - b)*(1+cnj c*c) = (a-c)*(1+cnj c*b) + (c-b)*(1 + cnj c*a)" by (simp add: field_simps) moreover have "1 + cnj c * c = 1 + (cmod c)\<^sup>2" using complex_norm_square by auto hence "cmod ((a - b)*(1+cnj c*c)) = cmod(a - b) * (1+(cmod c)\<^sup>2)" by (smt norm_mult norm_of_real zero_compare_simps(12)) ultimately have "cmod(a - b) * (1+(cmod c)\<^sup>2) \ cmod (a-c) * cmod (1+cnj c*b) + cmod (c-b) * cmod(1 + cnj c*a)" using complex_mod_triangle_ineq2[of "(a-c)*(1+cnj c*b)" "(c-b)*(1 + cnj c*a)"] - by simp + by (simp add: norm_mult) moreover have *: "\ a b c d b' d'. \b \ b'; d \ d'; a \ (0::real); c \ 0\ \ a*b + c*d \ a*b' + c*d'" - by (smt mult_left_mono) + by (simp add: add_mono_thms_linordered_semiring(1) mult_left_mono) have "cmod (a-c) * cmod (1+cnj c*b) + cmod (c-b) * cmod(1 + cnj c*a) \ cmod (a - c) * (sqrt (1+(cmod c)\<^sup>2) * sqrt (1+(cmod b)\<^sup>2)) + cmod (c - b) * (sqrt (1+(cmod c)\<^sup>2) * sqrt (1+(cmod a)\<^sup>2))" using *[OF cmod_1_plus_mult_le[of "cnj c" b] cmod_1_plus_mult_le[of "cnj c" a], of "cmod (a-c)" "cmod (c-b)"] by (simp add: field_simps real_sqrt_mult[symmetric]) ultimately have "cmod(a - b) * ?cc \ cmod (a - c) * sqrt ?cc * sqrt ?bb + cmod (c - b) * sqrt ?cc * sqrt ?aa" by simp moreover hence "0 \ ?cc * sqrt ?aa * sqrt ?bb" using mult_right_mono[of 0 "sqrt ?aa" "sqrt ?bb"] using mult_right_mono[of 0 "?cc" "sqrt ?aa * sqrt ?bb"] by simp moreover have "sqrt ?cc / ?cc = 1 / sqrt ?cc" using \sqrt ?cc > 0\ by (simp add: field_simps) hence "sqrt ?cc / (?cc * sqrt ?aa) = 1 / (sqrt ?aa * sqrt ?cc)" using times_divide_eq_right[of "1/sqrt ?aa" "sqrt ?cc" "?cc"] using \sqrt ?aa > 0\ by simp hence "cmod (a - c) * sqrt ?cc / (?cc * sqrt ?aa) = cmod (a - c) / (sqrt ?aa * sqrt ?cc)" using times_divide_eq_right[of "cmod (a - c)" "sqrt ?cc" "(?cc * sqrt ?aa)"] by simp moreover have "sqrt ?cc / ?cc = 1 / sqrt ?cc" using \sqrt ?cc > 0\ by (simp add: field_simps) hence "sqrt ?cc / (?cc * sqrt ?bb) = 1 / (sqrt ?bb * sqrt ?cc)" using times_divide_eq_right[of "1/sqrt ?bb" "sqrt ?cc" "?cc"] using \sqrt ?bb > 0\ by simp hence "cmod (c - b) * sqrt ?cc / (?cc * sqrt ?bb) = cmod (c - b) / (sqrt ?bb * sqrt ?cc)" using times_divide_eq_right[of "cmod (c - b)" "sqrt ?cc" "?cc * sqrt ?bb"] by simp ultimately show ?thesis using divide_right_mono[of "cmod (a - b) * ?cc" "cmod (a - c) * sqrt ?cc * sqrt ?bb + cmod (c - b) * sqrt ?cc * sqrt ?aa" "?cc * sqrt ?aa * sqrt ?bb"] \sqrt ?aa > 0\ \sqrt ?bb > 0\ \sqrt ?cc > 0\ by (simp add: add_divide_distrib) qed lemma dist_fs_triangle_infinite1: shows "1 / sqrt(1 + (cmod b)\<^sup>2) \ 1 / sqrt(1 + (cmod c)\<^sup>2) + cmod (b - c) / (sqrt(1 + (cmod b)\<^sup>2) * sqrt(1 + (cmod c)\<^sup>2))" proof- let ?bb = "sqrt (1 + (cmod b)\<^sup>2)" and ?cc = "sqrt (1 + (cmod c)\<^sup>2)" have "?bb > 0" "?cc > 0" by (metis add_strict_increasing real_sqrt_gt_0_iff zero_le_power2 zero_less_one)+ hence *: "?bb * ?cc \ 0" by simp have **: "(?cc - ?bb) / (?bb * ?cc) = 1 / ?bb - 1 / ?cc" using \sqrt (1 + (cmod b)\<^sup>2) > 0\ \sqrt (1 + (cmod c)\<^sup>2) > 0\ by (simp add: field_simps) show "1 / ?bb \ 1 / ?cc + cmod (b - c) / (?bb * ?cc)" using divide_right_mono[OF cmod_diff_ge[of c b] *] by (subst (asm) **) (simp add: field_simps norm_minus_commute) qed lemma dist_fs_triangle_infinite2: shows "1 / sqrt(1 + (cmod a)\<^sup>2) \ cmod (a - c) / (sqrt (1+(cmod a)\<^sup>2) * sqrt (1+(cmod c)\<^sup>2)) + 1 / sqrt(1 + (cmod c)\<^sup>2)" using dist_fs_triangle_infinite1[of a c] by simp lemma dist_fs_triangle_infinite3: shows "cmod(a - b) / (sqrt (1+(cmod a)\<^sup>2) * sqrt (1+(cmod b)\<^sup>2)) \ 1 / sqrt(1 + (cmod a)\<^sup>2) + 1 / sqrt(1 + (cmod b)\<^sup>2)" proof- let ?aa = "sqrt (1 + (cmod a)\<^sup>2)" and ?bb = "sqrt (1 + (cmod b)\<^sup>2)" have "?aa > 0" "?bb > 0" by (metis add_strict_increasing real_sqrt_gt_0_iff zero_le_power2 zero_less_one)+ hence *: "?aa * ?bb \ 0" by simp have **: "(?aa + ?bb) / (?aa * ?bb) = 1 / ?aa + 1 / ?bb" using \?aa > 0\ \?bb > 0\ by (simp add: field_simps) show "cmod (a - b) / (?aa * ?bb) \ 1 / ?aa + 1 / ?bb" using divide_right_mono[OF cmod_diff_le[of a b] *] by (subst (asm) **) (simp add: field_simps norm_minus_commute) qed lemma dist_fs_triangle: shows "dist_fs A B \ dist_fs A C + dist_fs C B" proof (cases "A = \\<^sub>h") case True show ?thesis proof (cases "B = \\<^sub>h") case True show ?thesis proof (cases "C = \\<^sub>h") case True show ?thesis using \A = \\<^sub>h\ \B = \\<^sub>h\ \C = \\<^sub>h\ by simp next case False then obtain c where "C = of_complex c" using inf_or_of_complex[of C] by auto show ?thesis using \A = \\<^sub>h\ \B = \\<^sub>h\ \C = of_complex c\ by (simp add: dist_fs_infinite2 dist_fs_sym) qed next case False then obtain b where "B = of_complex b" using inf_or_of_complex[of B] by auto show ?thesis proof (cases "C = \\<^sub>h") case True show ?thesis using \A = \\<^sub>h\ \C = \\<^sub>h\ \B = of_complex b\ by simp next case False then obtain c where "C = of_complex c" using inf_or_of_complex[of C] by auto show ?thesis using \A = \\<^sub>h\ \B = of_complex b\ \C = of_complex c\ using mult_left_mono[OF dist_fs_triangle_infinite1[of b c], of 2] by (simp add: dist_fs_finite dist_fs_infinite1 dist_fs_infinite2 dist_fs_sym) qed qed next case False then obtain a where "A = of_complex a" using inf_or_of_complex[of A] by auto show ?thesis proof (cases "B = \\<^sub>h") case True show ?thesis proof (cases "C = \\<^sub>h") case True show ?thesis using \B = \\<^sub>h\ \C = \\<^sub>h\ \A = of_complex a\ by (simp add: dist_fs_infinite2) next case False then obtain c where "C = of_complex c" using inf_or_of_complex[of C] by auto show ?thesis using \B = \\<^sub>h\ \C = of_complex c\ \A = of_complex a\ using mult_left_mono[OF dist_fs_triangle_infinite2[of a c], of 2] by (simp add: dist_fs_finite dist_fs_infinite1 dist_fs_infinite2) qed next case False then obtain b where "B = of_complex b" using inf_or_of_complex[of B] by auto show ?thesis proof (cases "C = \\<^sub>h") case True thus ?thesis using \C = \\<^sub>h\ \B = of_complex b\ \A = of_complex a\ using mult_left_mono[OF dist_fs_triangle_infinite3[of a b], of 2] by (simp add: dist_fs_finite dist_fs_infinite1 dist_fs_infinite2) next case False then obtain c where "C = of_complex c" using inf_or_of_complex[of C] by auto show ?thesis using \A = of_complex a\ \B = of_complex b\ \C = of_complex c\ using mult_left_mono[OF dist_fs_triangle_finite[of a b c], of 2] by (simp add: dist_fs_finite norm_minus_commute dist_fs_sym) qed qed qed (* -------------------------------------------------------------------------- *) subsubsection \$\mathbb{C}P^1$ with Fubini-Study metric is a metric space\ (* -------------------------------------------------------------------------- *) text \Using the (already available) fact that $\mathbb{R}^3$ is a metric space (under the distance function $\lambda\ x\ y.\ @{term norm}(x - y)$), it was not difficult to show that the type @{term complex_homo} equipped with @{term dist_fs} is a metric space, i.e., an instantiation of the @{term metric_space} locale.\ instantiation complex_homo :: metric_space begin definition "dist_complex_homo = dist_fs" definition "(uniformity_complex_homo :: (complex_homo \ complex_homo) filter) = (INF e\{0<..}. principal {(x, y). dist_class.dist x y < e})" definition "open_complex_homo (U :: complex_homo set) = (\ x \ U. eventually (\(x', y). x' = x \ y \ U) uniformity)" instance proof fix x y :: complex_homo show "(dist_class.dist x y = 0) = (x = y)" unfolding dist_complex_homo_def using dist_fs_zero1[of x] dist_fs_zero2[of x y] by auto next fix x y z :: complex_homo show "dist_class.dist x y \ dist_class.dist x z + dist_class.dist y z" unfolding dist_complex_homo_def using dist_fs_triangle[of x y z] by (simp add: dist_fs_sym) qed (simp_all add: open_complex_homo_def uniformity_complex_homo_def) end (* -------------------------------------------------------------------------- *) subsubsection \Chordal distance on the Riemann sphere\ (* -------------------------------------------------------------------------- *) text \Distance of the two points is given by the length of the chord. We show that it corresponds to the Fubini-Study metric in the plane.\ definition dist_riemann_sphere_r3 :: "R3 \ R3 \ real" where [simp]: "dist_riemann_sphere_r3 M1 M2 = (let (x1, y1, z1) = M1; (x2, y2, z2) = M2 in norm (x1 - x2, y1 - y2, z1 - z2))" lemma dist_riemann_sphere_r3_inner: assumes "M1 \ unit_sphere" and "M2 \ unit_sphere" shows "(dist_riemann_sphere_r3 M1 M2)\<^sup>2 = 2 - 2 * inner M1 M2" using assms apply (cases M1, cases M2) apply (auto simp add: norm_prod_def) apply (simp add: power2_eq_square field_simps) done lift_definition dist_riemann_sphere' :: "riemann_sphere \ riemann_sphere \ real" is dist_riemann_sphere_r3 done lemma dist_riemann_sphere_ge_0 [simp]: shows "dist_riemann_sphere' M1 M2 \ 0" apply transfer using norm_ge_zero by (simp add: split_def Let_def) text \Using stereographic projection we prove the connection between chordal metric on the spehere and Fubini-Study metric in the plane.\ lemma dist_stereographic_finite: assumes "stereographic M1 = of_complex m1" and "stereographic M2 = of_complex m2" shows "dist_riemann_sphere' M1 M2 = 2 * cmod (m1 - m2) / (sqrt (1 + (cmod m1)\<^sup>2) * sqrt (1 + (cmod m2)\<^sup>2))" using assms proof- have *: "M1 = inv_stereographic (of_complex m1)" "M2 = inv_stereographic (of_complex m2)" using inv_stereographic_is_inv assms by (metis inv_stereographic_stereographic)+ have "(1 + (cmod m1)\<^sup>2) \ 0" "(1 + (cmod m2)\<^sup>2) \ 0" by (smt power2_less_0)+ have "(1 + (cmod m1)\<^sup>2) > 0" "(1 + (cmod m2)\<^sup>2) > 0" by (smt realpow_square_minus_le)+ hence "(1 + (cmod m1)\<^sup>2) * (1 + (cmod m2)\<^sup>2) > 0" by (metis norm_mult_less norm_zero power2_eq_square zero_power2) hence ++: "sqrt ((1 + cmod m1 * cmod m1) * (1 + cmod m2 * cmod m2)) > 0" using real_sqrt_gt_0_iff by (simp add: power2_eq_square) hence **: "(2 * cmod (m1 - m2) / sqrt ((1 + cmod m1 * cmod m1) * (1 + cmod m2 * cmod m2))) \ 0 \ cmod (m1 - m2) \ 0" by (metis diff_self divide_nonneg_pos mult_2 norm_ge_zero norm_triangle_ineq4 norm_zero) have "(dist_riemann_sphere' M1 M2)\<^sup>2 * (1 + (cmod m1)\<^sup>2) * (1 + (cmod m2)\<^sup>2) = 4 * (cmod (m1 - m2))\<^sup>2" using * proof (transfer, transfer) fix m1 m2 M1 M2 assume us: "M1 \ unit_sphere" "M2 \ unit_sphere" and *: "M1 = inv_stereographic_cvec_r3 (of_complex_cvec m1)" "M2 = inv_stereographic_cvec_r3 (of_complex_cvec m2)" have "(1 + (cmod m1)\<^sup>2) \ 0" "(1 + (cmod m2)\<^sup>2) \ 0" by (smt power2_less_0)+ thus "(dist_riemann_sphere_r3 M1 M2)\<^sup>2 * (1 + (cmod m1)\<^sup>2) * (1 + (cmod m2)\<^sup>2) = 4 * (cmod (m1 - m2))\<^sup>2" apply (subst dist_riemann_sphere_r3_inner[OF us]) apply (subst *)+ apply (simp add: dist_riemann_sphere_r3_inner[OF us] complex_mult_cnj_cmod) apply (subst left_diff_distrib[of 2]) apply (subst left_diff_distrib[of "2*(1+(cmod m1)\<^sup>2)"]) apply (subst distrib_right[of _ _ "(1 + (cmod m1)\<^sup>2)"]) apply (subst distrib_right[of _ _ "(1 + (cmod m1)\<^sup>2)"]) apply simp apply (subst distrib_right[of _ _ "(1 + (cmod m2)\<^sup>2)"]) apply (subst distrib_right[of _ _ "(1 + (cmod m2)\<^sup>2)"]) apply (subst distrib_right[of _ _ "(1 + (cmod m2)\<^sup>2)"]) apply simp apply (subst (asm) cmod_square)+ apply (subst cmod_square)+ apply (simp add: field_simps) done qed hence "(dist_riemann_sphere' M1 M2)\<^sup>2 = 4 * (cmod (m1 - m2))\<^sup>2 / ((1 + (cmod m1)\<^sup>2) * (1 + (cmod m2)\<^sup>2))" using \(1 + (cmod m1)\<^sup>2) \ 0\ \(1 + (cmod m2)\<^sup>2) \ 0\ using eq_divide_imp[of "(1 + (cmod m1)\<^sup>2) * (1 + (cmod m2)\<^sup>2)" "(dist_riemann_sphere' M1 M2)\<^sup>2" "4 * (cmod (m1 - m2))\<^sup>2"] by simp thus "dist_riemann_sphere' M1 M2 = 2 * cmod (m1 - m2) / (sqrt (1 + (cmod m1)\<^sup>2) * sqrt (1 + (cmod m2)\<^sup>2))" using power2_eq_iff[of "dist_riemann_sphere' M1 M2" "2 * (cmod (m1 - m2)) / sqrt ((1 + (cmod m1)\<^sup>2) * (1 + (cmod m2)\<^sup>2))"] using \(1 + (cmod m1)\<^sup>2) * (1 + (cmod m2)\<^sup>2) > 0\ \(1 + (cmod m1)\<^sup>2) > 0\ \(1 + (cmod m2)\<^sup>2) > 0\ apply (auto simp add: power2_eq_square real_sqrt_mult[symmetric]) using dist_riemann_sphere_ge_0[of M1 M2] ** using ++ divide_le_0_iff by force qed lemma dist_stereographic_infinite: assumes "stereographic M1 = \\<^sub>h" and "stereographic M2 = of_complex m2" shows "dist_riemann_sphere' M1 M2 = 2 / sqrt (1 + (cmod m2)\<^sup>2)" proof- have *: "M1 = inv_stereographic \\<^sub>h" "M2 = inv_stereographic (of_complex m2)" using inv_stereographic_is_inv assms by (metis inv_stereographic_stereographic)+ have "(1 + (cmod m2)\<^sup>2) \ 0" by (smt power2_less_0) have "(1 + (cmod m2)\<^sup>2) > 0" by (smt realpow_square_minus_le)+ hence "sqrt (1 + cmod m2 * cmod m2) > 0" using real_sqrt_gt_0_iff by (simp add: power2_eq_square) hence **: "2 / sqrt (1 + cmod m2 * cmod m2) > 0" by simp have "(dist_riemann_sphere' M1 M2)\<^sup>2 * (1 + (cmod m2)\<^sup>2) = 4" using * apply transfer apply transfer proof- fix M1 M2 m2 assume us: "M1 \ unit_sphere" "M2 \ unit_sphere" and *: "M1 = inv_stereographic_cvec_r3 \\<^sub>v" "M2 = inv_stereographic_cvec_r3 (of_complex_cvec m2)" have "(1 + (cmod m2)\<^sup>2) \ 0" by (smt power2_less_0) thus "(dist_riemann_sphere_r3 M1 M2)\<^sup>2 * (1 + (cmod m2)\<^sup>2) = 4" apply (subst dist_riemann_sphere_r3_inner[OF us]) apply (subst *)+ apply (simp add: complex_mult_cnj_cmod) apply (subst left_diff_distrib[of 2], simp) done qed hence "(dist_riemann_sphere' M1 M2)\<^sup>2 = 4 / (1 + (cmod m2)\<^sup>2)" using \(1 + (cmod m2)\<^sup>2) \ 0\ by (simp add: field_simps) thus "dist_riemann_sphere' M1 M2 = 2 / sqrt (1 + (cmod m2)\<^sup>2)" using power2_eq_iff[of "dist_riemann_sphere' M1 M2" "2 / sqrt (1 + (cmod m2)\<^sup>2)"] using \(1 + (cmod m2)\<^sup>2) > 0\ apply (auto simp add: power2_eq_square real_sqrt_mult[symmetric]) using dist_riemann_sphere_ge_0[of M1 M2] ** by simp qed lemma dist_rieman_sphere_zero [simp]: shows "dist_riemann_sphere' M M = 0" by transfer auto lemma dist_riemann_sphere_sym: shows "dist_riemann_sphere' M1 M2 = dist_riemann_sphere' M2 M1" proof transfer fix M1 M2 :: R3 obtain x1 y1 z1 x2 y2 z2 where MM: "(x1, y1, z1) = M1" "(x2, y2, z2) = M2" by (cases "M1", cases "M2", auto) show "dist_riemann_sphere_r3 M1 M2 = dist_riemann_sphere_r3 M2 M1" using norm_minus_cancel[of "(x1 - x2, y1 - y2, z1 - z2)"] MM[symmetric] by simp qed text \Central theorem that connects the two metrics.\ lemma dist_stereographic: shows "dist_riemann_sphere' M1 M2 = dist_fs (stereographic M1) (stereographic M2)" proof (cases "M1 = North") case True hence "stereographic M1 = \\<^sub>h" by (simp add: stereographic_North) show ?thesis proof (cases "M2 = North") case True show ?thesis using \M1 = North\ \M2 = North\ by auto next case False hence "stereographic M2 \ \\<^sub>h" using stereographic_North[of M2] by simp then obtain m2 where "stereographic M2 = of_complex m2" using inf_or_of_complex[of "stereographic M2"] by auto show ?thesis using \stereographic M2 = of_complex m2\ \stereographic M1 = \\<^sub>h\ using dist_fs_infinite1 dist_stereographic_infinite by (simp add: dist_fs_sym) qed next case False hence "stereographic M1 \ \\<^sub>h" by (simp add: stereographic_North) then obtain m1 where "stereographic M1 = of_complex m1" using inf_or_of_complex[of "stereographic M1"] by auto show ?thesis proof (cases "M2 = North") case True hence "stereographic M2 = \\<^sub>h" by (simp add: stereographic_North) show ?thesis using \stereographic M1 = of_complex m1\ \stereographic M2 = \\<^sub>h\ using dist_fs_infinite2 dist_stereographic_infinite by (subst dist_riemann_sphere_sym, simp add: dist_fs_sym) next case False hence "stereographic M2 \ \\<^sub>h" by (simp add: stereographic_North) then obtain m2 where "stereographic M2 = of_complex m2" using inf_or_of_complex[of "stereographic M2"] by auto show ?thesis using \stereographic M1 = of_complex m1\ \stereographic M2 = of_complex m2\ using dist_fs_finite dist_stereographic_finite by simp qed qed text \Other direction\ lemma dist_stereographic': shows "dist_fs A B = dist_riemann_sphere' (inv_stereographic A) (inv_stereographic B)" by (subst dist_stereographic) (metis stereographic_inv_stereographic) text \The @{term riemann_sphere} equipped with @{term dist_riemann_sphere'} is a metric space, i.e., an instantiation of the @{term metric_space} locale.\ instantiation riemann_sphere :: metric_space begin definition "dist_riemann_sphere = dist_riemann_sphere'" definition "(uniformity_riemann_sphere :: (riemann_sphere \ riemann_sphere) filter) = (INF e\{0<..}. principal {(x, y). dist_class.dist x y < e})" definition "open_riemann_sphere (U :: riemann_sphere set) = (\ x \ U. eventually (\(x', y). x' = x \ y \ U) uniformity)" instance proof fix x y :: riemann_sphere show "(dist_class.dist x y = 0) = (x = y)" unfolding dist_riemann_sphere_def proof transfer fix x y :: R3 obtain x1 y1 z1 x2 y2 z2 where *: "(x1, y1, z1) = x" "(x2, y2, z2) = y" by (cases x, cases y, auto) assume "x \ unit_sphere" "y \ unit_sphere" thus "(dist_riemann_sphere_r3 x y = 0) = (x = y)" using norm_eq_zero[of "(x1 - y2, y1 - y2, z1 - z2)"] using *[symmetric] by (simp add: zero_prod_def) qed next fix x y z :: riemann_sphere show "dist_class.dist x y \ dist_class.dist x z + dist_class.dist y z" unfolding dist_riemann_sphere_def proof transfer fix x y z :: R3 obtain x1 y1 z1 x2 y2 z2 x3 y3 z3 where MM: "(x1, y1, z1) = x" "(x2, y2, z2) = y" "(x3, y3, z3) = z" by (cases "x", cases "y", cases "z", auto) assume "x \ unit_sphere" "y \ unit_sphere" "z \ unit_sphere" thus "dist_riemann_sphere_r3 x y \ dist_riemann_sphere_r3 x z + dist_riemann_sphere_r3 y z" using MM[symmetric] norm_minus_cancel[of "(x3 - x2, y3 - y2, z3 - z2)"] norm_triangle_ineq[of "(x1 - x3, y1 - y3, z1 - z3)" "(x3 - x2, y3 - y2, z3 - z2)"] by simp qed qed (simp_all add: uniformity_riemann_sphere_def open_riemann_sphere_def) end text \The @{term riemann_sphere} metric space is perfect, i.e., it does not have isolated points.\ instantiation riemann_sphere :: perfect_space begin instance proof fix M :: riemann_sphere show "\ open {M}" unfolding open_dist dist_riemann_sphere_def apply (subst dist_riemann_sphere_sym) proof transfer fix M assume "M \ unit_sphere" obtain x y z where MM: "M = (x, y, z)" by (cases "M") auto then obtain \ \ where *: "x = cos \ * cos \" "y = cos \ * sin \" "z = sin \" "-pi / 2 \ \ \ \ \ pi / 2" using \M \ unit_sphere\ using ex_sphere_params[of x y z] by auto have "\ e. e > 0 \ (\y. y \ unit_sphere \ dist_riemann_sphere_r3 M y < e \ y \ M)" proof- fix e :: real assume "e > 0" then obtain \' where "1 - (e*e/2) < cos (\ - \')" "\ \ \'" "-pi/2 \ \'" "\' \ pi/2" using ex_cos_gt[of \ "1 - (e*e/2)"] \- pi / 2 \ \ \ \ \ pi / 2\ by auto hence "sin \ \ sin \'" using \-pi / 2 \ \ \ \ \ pi / 2\ sin_inj[of \ \'] by auto have "2 - 2 * cos (\ - \') < e*e" using mult_strict_right_mono[OF \1 - (e*e/2) < cos (\ - \')\, of 2] by (simp add: field_simps) have "2 - 2 * cos (\ - \') \ 0" using cos_le_one[of "\ - \'"] by (simp add: algebra_split_simps) let ?M' = "(cos \' * cos \, cos \' * sin \, sin \')" have "dist_riemann_sphere_r3 M ?M' = sqrt ((cos \ - cos \')\<^sup>2 + (sin \ - sin \')\<^sup>2)" using MM * sphere_params_on_sphere[of _ \' \] using sin_cos_squared_add[of \] apply (simp add: dist_riemann_sphere'_def Abs_riemann_sphere_inverse norm_prod_def) apply (subst left_diff_distrib[symmetric])+ apply (subst power_mult_distrib)+ apply (subst distrib_left[symmetric]) apply simp done also have "... = sqrt (2 - 2*cos (\ - \'))" by (simp add: power2_eq_square field_simps cos_diff) finally have "(dist_riemann_sphere_r3 M ?M')\<^sup>2 = 2 - 2*cos (\ - \')" using \2 - 2 * cos (\ - \') \ 0\ by simp hence "(dist_riemann_sphere_r3 M ?M')\<^sup>2 < e\<^sup>2" using \2 - 2 * cos (\ - \') < e*e\ by (simp add: power2_eq_square) hence "dist_riemann_sphere_r3 M ?M' < e" apply (rule power2_less_imp_less) using \e > 0\ by simp moreover have "M \ ?M'" using MM \sin \ \ sin \'\ * by simp moreover have "?M' \ unit_sphere" using sphere_params_on_sphere by auto ultimately show "\y. y \ unit_sphere \ dist_riemann_sphere_r3 M y < e \ y \ M" unfolding dist_riemann_sphere_def by (rule_tac x="?M'" in exI, simp) qed thus "\ (\x\{M}. \e>0. \y\{x. x \ unit_sphere}. dist_riemann_sphere_r3 x y < e \ y \ {M})" by auto qed qed end text \The @{term complex_homo} metric space is perfect, i.e., it does not have isolated points.\ instantiation complex_homo :: perfect_space begin instance proof fix x::complex_homo show "\ open {x}" unfolding open_dist proof (auto) fix e::real assume "e > 0" thus "\ y. dist_class.dist y x < e \ y \ x" using not_open_singleton[of "inv_stereographic x"] unfolding open_dist unfolding dist_complex_homo_def dist_riemann_sphere_def apply (subst dist_stereographic', auto) apply (erule_tac x=e in allE, auto) apply (rule_tac x="stereographic y" in exI, auto) done qed qed end lemma Lim_within: shows "(f \ l) (at a within S) \ (\e >0. \d>0. \x \ S. 0 < dist_class.dist x a \ dist_class.dist x a < d \ dist_class.dist (f x) l < e)" by (auto simp: tendsto_iff eventually_at) lemma continuous_on_iff: shows "continuous_on s f \ (\x\s. \e>0. \d>0. \x'\s. dist_class.dist x' x < d \ dist_class.dist (f x') (f x) < e)" unfolding continuous_on_def Lim_within by (metis dist_pos_lt dist_self) text \Using the chordal metric in the extended plane, and the Euclidean metric on the sphere in $\mathbb{R}^3$, the stereographic and inverse stereographic projections are proved to be continuous.\ lemma "continuous_on UNIV stereographic" unfolding continuous_on_iff unfolding dist_complex_homo_def dist_riemann_sphere_def by (subst dist_stereographic', auto) lemma "continuous_on UNIV inv_stereographic" unfolding continuous_on_iff unfolding dist_complex_homo_def dist_riemann_sphere_def by (subst dist_stereographic, auto) (* -------------------------------------------------------------------------- *) subsubsection \Chordal circles\ (* -------------------------------------------------------------------------- *) text \Real circlines are sets of points that are equidistant from some given point in the chordal metric. There are exactly two such points (two chordal centers). On the Riemann sphere, these two points are obtained as intersections of the sphere and a line that goes trough center of the circle, and its orthogonal to its plane.\ text \The circline for the given chordal center and radius.\ definition chordal_circle_cvec_cmat :: "complex_vec \ real \ complex_mat" where [simp]: "chordal_circle_cvec_cmat a r = (let (a1, a2) = a in ((4*a2*cnj a2 - (cor r)\<^sup>2*(a1*cnj a1 + a2*cnj a2)), (-4*a1*cnj a2), (-4*cnj a1*a2), (4*a1*cnj a1 - (cor r)\<^sup>2*(a1*cnj a1 + a2*cnj a2))))" lemma chordal_circle_cmat_hermitean_nonzero [simp]: assumes "a \ vec_zero" shows "chordal_circle_cvec_cmat a r \ hermitean_nonzero" using assms by (cases a) (auto simp add: hermitean_def mat_adj_def mat_cnj_def Let_def) lift_definition chordal_circle_hcoords_clmat :: "complex_homo_coords \ real \ circline_mat" is chordal_circle_cvec_cmat using chordal_circle_cmat_hermitean_nonzero by (simp del: chordal_circle_cvec_cmat_def) lift_definition chordal_circle :: "complex_homo \ real \ circline" is chordal_circle_hcoords_clmat proof transfer fix a b :: complex_vec and r :: real assume *: "a \ vec_zero" "b \ vec_zero" obtain a1 a2 where aa: "a = (a1, a2)" by (cases a, auto) obtain b1 b2 where bb: "b = (b1, b2)" by (cases b, auto) assume "a \\<^sub>v b" then obtain k where "b = (k * a1, k * a2)" "k \ 0" using aa bb by auto moreover have "cor (Re (k * cnj k)) = k * cnj k" by (metis complex_In_mult_cnj_zero complex_of_real_Re) ultimately show "circline_eq_cmat (chordal_circle_cvec_cmat a r) (chordal_circle_cvec_cmat b r)" using * aa bb by simp (rule_tac x="Re (k*cnj k)" in exI, auto simp add: Let_def field_simps) qed lemma sqrt_1_plus_square: shows "sqrt (1 + a\<^sup>2) \ 0" by (smt real_sqrt_less_mono real_sqrt_zero realpow_square_minus_le) lemma assumes "dist_fs z a = r" shows "z \ circline_set (chordal_circle a r)" proof (cases "a \ \\<^sub>h") case True then obtain a' where "a = of_complex a'" using inf_or_of_complex by auto let ?A = "4 - (cor r)\<^sup>2 * (1 + (a'*cnj a'))" and ?B = "-4*a'" and ?C="-4*cnj a'" and ?D = "4*a'*cnj a' - (cor r)\<^sup>2 * (1 + (a'*cnj a'))" have hh: "(?A, ?B, ?C, ?D) \ {H. hermitean H \ H \ mat_zero}" by (auto simp add: hermitean_def mat_adj_def mat_cnj_def power2_eq_square) hence *: "chordal_circle (of_complex a') r = mk_circline ?A ?B ?C ?D" by (transfer, transfer, simp, rule_tac x=1 in exI, simp) show ?thesis proof (cases "z \ \\<^sub>h") case True then obtain z' where "z = of_complex z'" using inf_or_of_complex[of z] inf_or_of_complex[of a] by auto have "2 * cmod (z' - a') / (sqrt (1 + (cmod z')\<^sup>2) * sqrt (1 + (cmod a')\<^sup>2)) = r" using dist_fs_finite[of z' a'] assms \z = of_complex z'\ \a = of_complex a'\ by auto hence "4 * (cmod (z' - a'))\<^sup>2 / ((1 + (cmod z')\<^sup>2) * (1 + (cmod a')\<^sup>2)) = r\<^sup>2 " by (auto simp add: field_simps) moreover have "sqrt (1 + (cmod z')\<^sup>2) \ 0" "sqrt (1 + (cmod a')\<^sup>2) \ 0" using sqrt_1_plus_square by simp+ hence "(1 + (cmod z')\<^sup>2) * (1 + (cmod a')\<^sup>2) \ 0" by simp ultimately have "4 * (cmod (z' - a'))\<^sup>2 = r\<^sup>2 * ((1 + (cmod z')\<^sup>2) * (1 + (cmod a')\<^sup>2))" by (simp add: field_simps) hence "4 * Re ((z' - a')*cnj (z' - a')) = r\<^sup>2 * (1 + Re (z'*cnj z')) * (1 + Re (a'*cnj a'))" by ((subst cmod_square[symmetric])+, simp) hence "4 * (Re(z'*cnj z') - Re(a'*cnj z') - Re(cnj a'*z') + Re(a'*cnj a')) = r\<^sup>2 * (1 + Re (z'*cnj z')) * (1 + Re (a'*cnj a'))" by (simp add: field_simps) hence "Re (?A * z' * cnj z' + ?B * cnj z' + ?C * z' + ?D) = 0" by (simp add: power2_eq_square field_simps) hence "?A * z' * cnj z' + ?B * cnj z' + ?C * z' + ?D = 0" by (subst complex_eq_if_Re_eq) (auto simp add: power2_eq_square) hence "(cnj z' * ?A + ?C) * z' + (cnj z' * ?B + ?D) = 0" by algebra hence "z \ circline_set (mk_circline ?A ?B ?C ?D)" using \z = of_complex z'\ hh unfolding circline_set_def by simp (transfer, transfer, simp add: vec_cnj_def) thus ?thesis using * by (subst \a = of_complex a'\) simp next case False hence "2 / sqrt (1 + (cmod a')\<^sup>2) = r" using assms \a = of_complex a'\ using dist_fs_infinite2[of a'] by simp moreover have "sqrt (1 + (cmod a')\<^sup>2) \ 0" using sqrt_1_plus_square by simp ultimately have "2 = r * sqrt (1 + (cmod a')\<^sup>2)" by (simp add: field_simps) hence "4 = (r * sqrt (1 + (cmod a')\<^sup>2))\<^sup>2" by simp hence "4 = r\<^sup>2 * (1 + (cmod a')\<^sup>2)" by (simp add: power_mult_distrib) hence "Re (4 - (cor r)\<^sup>2 * (1 + (a' * cnj a'))) = 0" by (subst (asm) cmod_square) (simp add: field_simps power2_eq_square) hence "4 - (cor r)\<^sup>2 * (1 + (a'*cnj a')) = 0" by (subst complex_eq_if_Re_eq) (auto simp add: power2_eq_square) hence "circline_A0 (mk_circline ?A ?B ?C ?D)" using hh by (simp, transfer, transfer, simp) hence "z \ circline_set (mk_circline ?A ?B ?C ?D)" using inf_in_circline_set[of "mk_circline ?A ?B ?C ?D"] using \\ z \ \\<^sub>h\ by simp thus ?thesis using * by (subst \a = of_complex a'\) simp qed next case False let ?A = "-(cor r)\<^sup>2" and ?B = "0" and ?C = "0" and ?D = "4 -(cor r)\<^sup>2" have hh: "(?A, ?B, ?C, ?D) \ {H. hermitean H \ H \ mat_zero}" by (auto simp add: hermitean_def mat_adj_def mat_cnj_def power2_eq_square) hence *: "chordal_circle a r = mk_circline ?A ?B ?C ?D" using \\ a \ \\<^sub>h\ by simp (transfer, transfer, simp, rule_tac x=1 in exI, simp) show ?thesis proof (cases "z = \\<^sub>h") case True show ?thesis using assms \z = \\<^sub>h\ \\ a \ \\<^sub>h\ using * hh by (simp, subst inf_in_circline_set, transfer, transfer, simp) next case False then obtain z' where "z = of_complex z'" using inf_or_of_complex[of z] by auto have "2 / sqrt (1 + (cmod z')\<^sup>2) = r" using assms \z = of_complex z'\\\ a \ \\<^sub>h\ using dist_fs_infinite2[of z'] by (simp add: dist_fs_sym) moreover have "sqrt (1 + (cmod z')\<^sup>2) \ 0" using sqrt_1_plus_square by simp ultimately have "2 = r * sqrt (1 + (cmod z')\<^sup>2)" by (simp add: field_simps) hence "4 = (r * sqrt (1 + (cmod z')\<^sup>2))\<^sup>2" by simp hence "4 = r\<^sup>2 * (1 + (cmod z')\<^sup>2)" by (simp add: power_mult_distrib) hence "Re (4 - (cor r)\<^sup>2 * (1 + (z' * cnj z'))) = 0" by (subst (asm) cmod_square) (simp add: field_simps power2_eq_square) hence "- (cor r)\<^sup>2 * z'*cnj z' + 4 - (cor r)\<^sup>2 = 0" by (subst complex_eq_if_Re_eq) (auto simp add: power2_eq_square field_simps) hence "z \ circline_set (mk_circline ?A ?B ?C ?D)" using hh unfolding circline_set_def by (subst \z = of_complex z'\, simp) (transfer, transfer, auto simp add: vec_cnj_def field_simps) thus ?thesis using * by simp qed qed lemma assumes "z \ circline_set (chordal_circle a r)" and "r \ 0" shows "dist_fs z a = r" proof (cases "a = \\<^sub>h") case False then obtain a' where "a = of_complex a'" using inf_or_of_complex by auto let ?A = "4 - (cor r)\<^sup>2 * (1 + (a'*cnj a'))" and ?B = "-4*a'" and ?C="-4*cnj a'" and ?D = "4*a'*cnj a' - (cor r)\<^sup>2 * (1 + (a'*cnj a'))" have hh: "(?A, ?B, ?C, ?D) \ {H. hermitean H \ H \ mat_zero}" by (auto simp add: hermitean_def mat_adj_def mat_cnj_def power2_eq_square) hence *: "chordal_circle (of_complex a') r = mk_circline ?A ?B ?C ?D" by (transfer, transfer, simp, rule_tac x=1 in exI, simp) show ?thesis proof (cases "z = \\<^sub>h") case False then obtain z' where "z = of_complex z'" using inf_or_of_complex[of z] inf_or_of_complex[of a] by auto hence "z \ circline_set (mk_circline ?A ?B ?C ?D)" using assms \a = of_complex a'\ * by simp hence "(cnj z' * ?A + ?C) * z' + (cnj z' * ?B + ?D) = 0" using hh unfolding circline_set_def by (subst (asm) \z = of_complex z'\, simp) (transfer, transfer, simp add: vec_cnj_def) hence "?A * z' * cnj z' + ?B * cnj z' + ?C * z' + ?D = 0" by algebra hence "Re (?A * z' * cnj z' + ?B * cnj z' +?C * z' +?D) = 0" by (simp add: power2_eq_square field_simps) hence "4 * Re ((z' - a')*cnj (z' - a')) = r\<^sup>2 * (1 + Re (z'*cnj z')) * (1 + Re (a'*cnj a'))" by (simp add: field_simps power2_eq_square) hence "4 * (cmod (z' - a'))\<^sup>2 = r\<^sup>2 * ((1 + (cmod z')\<^sup>2) * (1 + (cmod a')\<^sup>2))" by (subst cmod_square)+ simp moreover have "sqrt (1 + (cmod z')\<^sup>2) \ 0" "sqrt (1 + (cmod a')\<^sup>2) \ 0" using sqrt_1_plus_square by simp+ hence "(1 + (cmod z')\<^sup>2) * (1 + (cmod a')\<^sup>2) \ 0" by simp ultimately have "4 * (cmod (z' - a'))\<^sup>2 / ((1 + (cmod z')\<^sup>2) * (1 + (cmod a')\<^sup>2)) = r\<^sup>2 " by (simp add: field_simps) hence "2 * cmod (z' - a') / (sqrt (1 + (cmod z')\<^sup>2) * sqrt (1 + (cmod a')\<^sup>2)) = r" using \r \ 0\ by (subst (asm) real_sqrt_eq_iff[symmetric]) (simp add: real_sqrt_mult real_sqrt_divide) thus ?thesis using \z = of_complex z'\ \a = of_complex a'\ using dist_fs_finite[of z' a'] by simp next case True have "z \ circline_set (mk_circline ?A ?B ?C ?D)" using assms \a = of_complex a'\ * by simp hence "circline_A0 (mk_circline ?A ?B ?C ?D)" using inf_in_circline_set[of "mk_circline ?A ?B ?C ?D"] using \z = \\<^sub>h\ by simp hence "4 - (cor r)\<^sup>2 * (1 + (a'*cnj a')) = 0" using hh by (transfer, transfer, simp) hence "Re (4 - (cor r)\<^sup>2 * (1 + (a' * cnj a'))) = 0" by simp hence "4 = r\<^sup>2 * (1 + (cmod a')\<^sup>2)" by (subst cmod_square) (simp add: power2_eq_square) hence "2 = r * sqrt (1 + (cmod a')\<^sup>2)" using \r \ 0\ by (subst (asm) real_sqrt_eq_iff[symmetric]) (simp add: real_sqrt_mult) moreover have "sqrt (1 + (cmod a')\<^sup>2) \ 0" using sqrt_1_plus_square by simp ultimately have "2 / sqrt (1 + (cmod a')\<^sup>2) = r" by (simp add: field_simps) thus ?thesis using \a = of_complex a'\ \z = \\<^sub>h\ using dist_fs_infinite2[of a'] by simp qed next case True let ?A = "-(cor r)\<^sup>2" and ?B = "0" and ?C = "0" and ?D = "4 -(cor r)\<^sup>2" have hh: "(?A, ?B, ?C, ?D) \ {H. hermitean H \ H \ mat_zero}" by (auto simp add: hermitean_def mat_adj_def mat_cnj_def power2_eq_square) hence *: "chordal_circle a r = mk_circline ?A ?B ?C ?D" using \a = \\<^sub>h\ by simp (transfer, transfer, simp, rule_tac x=1 in exI, simp) show ?thesis proof (cases "z = \\<^sub>h") case True thus ?thesis using \a = \\<^sub>h\ assms * hh by simp (subst (asm) inf_in_circline_set, transfer, transfer, simp) next case False then obtain z' where "z = of_complex z'" using inf_or_of_complex by auto hence "z \ circline_set (mk_circline ?A ?B ?C ?D)" using assms * by simp hence "- (cor r)\<^sup>2 * z'*cnj z' + 4 - (cor r)\<^sup>2 = 0" using hh unfolding circline_set_def apply (subst (asm) \z = of_complex z'\) by (simp, transfer, transfer, simp add: vec_cnj_def, algebra) hence "4 - (cor r)\<^sup>2 * (1 + (z'*cnj z')) = 0" by (simp add: field_simps) hence "Re (4 - (cor r)\<^sup>2 * (1 + (z' * cnj z'))) = 0" by simp hence "4 = r\<^sup>2 * (1 + (cmod z')\<^sup>2)" by (subst cmod_square) (simp add: power2_eq_square) hence "2 = r * sqrt (1 + (cmod z')\<^sup>2)" using \r \ 0\ by (subst (asm) real_sqrt_eq_iff[symmetric]) (simp add: real_sqrt_mult) moreover have "sqrt (1 + (cmod z')\<^sup>2) \ 0" using sqrt_1_plus_square by simp ultimately have "2 / sqrt (1 + (cmod z')\<^sup>2) = r" by (simp add: field_simps) thus ?thesis using \z = of_complex z'\ \a = \\<^sub>h\ using dist_fs_infinite2[of z'] by (simp add: dist_fs_sym) qed qed text \Two chordal centers and radii for the given circline\ definition chordal_circles_cmat :: "complex_mat \ (complex \ real) \ (complex \ real)" where [simp]: "chordal_circles_cmat H = (let (A, B, C, D) = H; dsc = sqrt(Re ((D-A)\<^sup>2 + 4 * (B*cnj B))); a1 = (A - D + cor dsc) / (2 * C); r1 = sqrt((4 - Re((-4 * a1/B) * A)) / (1 + Re (a1*cnj a1))); a2 = (A - D - cor dsc) / (2 * C); r2 = sqrt((4 - Re((-4 * a2/B) * A)) / (1 + Re (a2*cnj a2))) in ((a1, r1), (a2, r2)))" lift_definition chordal_circles_clmat :: "circline_mat \ (complex \ real) \ (complex \ real)" is chordal_circles_cmat done lift_definition chordal_circles :: "ocircline \ (complex \ real) \ (complex \ real)" is chordal_circles_clmat proof transfer fix H1 H2 :: complex_mat obtain A1 B1 C1 D1 where hh1: "(A1, B1, C1, D1) = H1" by (cases H1) auto obtain A2 B2 C2 D2 where hh2: "(A2, B2, C2, D2) = H2" by (cases H2) auto assume "ocircline_eq_cmat H1 H2" then obtain k where *: "k > 0" "A2 = cor k * A1" "B2 = cor k * B1" "C2 = cor k * C1" "D2 = cor k * D1" using hh1[symmetric] hh2[symmetric] by auto let ?dsc1 = "sqrt (Re ((D1 - A1)\<^sup>2 + 4 * (B1 * cnj B1)))" and ?dsc2 = "sqrt (Re ((D2 - A2)\<^sup>2 + 4 * (B2 * cnj B2)))" let ?a11 = "(A1 - D1 + cor ?dsc1) / (2 * C1)" and ?a12 = "(A2 - D2 + cor ?dsc2) / (2 * C2)" let ?a21 = "(A1 - D1 - cor ?dsc1) / (2 * C1)" and ?a22 = "(A2 - D2 - cor ?dsc2) / (2 * C2)" let ?r11 = "sqrt((4 - Re((-4 * ?a11/B1) * A1)) / (1 + Re (?a11*cnj ?a11)))" let ?r12 = "sqrt((4 - Re((-4 * ?a12/B2) * A2)) / (1 + Re (?a12*cnj ?a12)))" let ?r21 = "sqrt((4 - Re((-4 * ?a21/B1) * A1)) / (1 + Re (?a21*cnj ?a21)))" let ?r22 = "sqrt((4 - Re((-4 * ?a22/B2) * A2)) / (1 + Re (?a22*cnj ?a22)))" have "Re ((D2 - A2)\<^sup>2 + 4 * (B2 * cnj B2)) = k\<^sup>2 * Re ((D1 - A1)\<^sup>2 + 4 * (B1 * cnj B1))" using * by (simp add: power2_eq_square field_simps) hence "?dsc2 = k * ?dsc1" using \k > 0\ by (simp add: real_sqrt_mult) hence "A2 - D2 + cor ?dsc2 = cor k * (A1 - D1 + cor ?dsc1)" "A2 - D2 - cor ?dsc2 = cor k * (A1 - D1 - cor ?dsc1)" "2*C2 = cor k * (2*C1)" using * by (auto simp add: field_simps) hence "?a12 = ?a11" "?a22 = ?a21" using \k > 0\ by simp_all moreover have "Re((-4 * ?a12/B2) * A2) = Re((-4 * ?a11/B1) * A1)" using * by (subst \?a12 = ?a11\) (simp, simp add: field_simps) have "?r12 = ?r11" by (subst \Re((-4 * ?a12/B2) * A2) = Re((-4 * ?a11/B1) * A1)\, (subst \?a12 = ?a11\)+) simp moreover have "Re((-4 * ?a22/B2) * A2) = Re((-4 * ?a21/B1) * A1)" using * by (subst \?a22 = ?a21\) (simp, simp add: field_simps) have "?r22 = ?r21" by (subst \Re((-4 * ?a22/B2) * A2) = Re((-4 * ?a21/B1) * A1)\, (subst \?a22 = ?a21\)+) simp moreover have "chordal_circles_cmat H1 = ((?a11, ?r11), (?a21, ?r21))" using hh1[symmetric] unfolding chordal_circles_cmat_def Let_def by (simp del: times_complex.sel) moreover have "chordal_circles_cmat H2 = ((?a12, ?r12), (?a22, ?r22))" using hh2[symmetric] unfolding chordal_circles_cmat_def Let_def by (simp del: times_complex.sel) ultimately show "chordal_circles_cmat H1 = chordal_circles_cmat H2" by metis qed lemma chordal_circle_radius_positive: assumes "hermitean (A, B, C, D)" and "Re (mat_det (A, B, C, D)) \ 0" and "B \ 0" and "dsc = sqrt(Re ((D-A)\<^sup>2 + 4 * (B*cnj B)))" and "a1 = (A - D + cor dsc) / (2 * C)" and "a2 = (A - D - cor dsc) / (2 * C)" shows "Re (A*a1/B) \ -1 \ Re (A*a2/B) \ -1" proof- from assms have "is_real A" "is_real D" "C = cnj B" using hermitean_elems by auto have *: "A*a1/B = ((A - D + cor dsc) / (2 * (B * cnj B))) * A" using \B \ 0\ \C = cnj B\ \a1 = (A - D + cor dsc) / (2 * C)\ by (simp add: field_simps) algebra have **: "A*a2/B = ((A - D - cor dsc) / (2 * (B * cnj B))) * A" using \B \ 0\ \C = cnj B\ \a2 = (A - D - cor dsc) / (2 * C)\ by (simp add: field_simps) algebra have "dsc \ 0" proof- have "0 \ Re ((D - A)\<^sup>2) + 4 * Re ((cor (cmod B))\<^sup>2)" - using \is_real A\ \is_real D\ - by (subst cor_squared, subst Re_complex_of_real) (simp add: power2_eq_square) + using \is_real A\ \is_real D\ by simp thus ?thesis using \dsc = sqrt(Re ((D-A)\<^sup>2 + 4*(B*cnj B)))\ by (subst (asm) complex_mult_cnj_cmod) simp qed hence "Re (A - D - cor dsc) \ Re (A - D + cor dsc)" by simp moreover have "Re (2 * (B * cnj B)) > 0" using \B \ 0\ by (subst complex_mult_cnj_cmod, simp add: power2_eq_square) ultimately have xxx: "Re (A - D + cor dsc) / Re (2 * (B * cnj B)) \ Re (A - D - cor dsc) / Re (2 * (B * cnj B))" (is "?lhs \ ?rhs") by (metis divide_right_mono less_eq_real_def) have "Re A * Re D \ Re (B*cnj B)" using \Re (mat_det (A, B, C, D)) \ 0\ \C = cnj B\ \is_real A\ \is_real D\ by simp have "(Re (2 * (B * cnj B)) / Re A) / Re (2 * B * cnj B) = 1 / Re A" using \Re (2 * (B * cnj B)) > 0\ apply (subst divide_divide_eq_left) apply (subst mult.assoc) apply (subst nonzero_divide_mult_cancel_right) by simp_all show ?thesis proof (cases "Re A > 0") case True hence "Re (A*a1/B) \ Re (A*a2/B)" using * ** \Re (2 * (B * cnj B)) > 0\ \B \ 0\ \is_real A\ \is_real D\ xxx using mult_right_mono[of ?rhs ?lhs "Re A"] apply simp apply (subst Re_divide_real, simp, simp) apply (subst Re_divide_real, simp, simp) apply (subst Re_mult_real, simp)+ apply simp done moreover have "Re (A*a2/B) \ -1" proof- from \Re A * Re D \ Re (B*cnj B)\ have "Re (A\<^sup>2) \ Re (B*cnj B) + Re ((A - D)*A)" using \Re A > 0\ \is_real A\ \is_real D\ by (simp add: power2_eq_square field_simps) have "1 \ Re (B*cnj B) / Re (A\<^sup>2) + Re (A - D) / Re A" using \Re A > 0\ \is_real A\ \is_real D\ using divide_right_mono[OF \Re (A\<^sup>2) \ Re (B*cnj B) + Re ((A - D)*A)\, of "Re (A\<^sup>2)"] by (simp add: power2_eq_square add_divide_distrib) have "4 * Re(B*cnj B) \ 4 * (Re (B*cnj B))\<^sup>2 / Re (A\<^sup>2) + 2*Re (A - D) / Re A * 2 * Re(B*cnj B)" using mult_right_mono[OF \1 \ Re (B*cnj B) / Re (A\<^sup>2) + Re (A - D) / Re A\, of "4 * Re (B*cnj B)"] by (simp add: distrib_right) (simp add: power2_eq_square field_simps) moreover have "A \ 0" using \Re A > 0\ by auto hence "4 * (Re (B*cnj B))\<^sup>2 / Re (A\<^sup>2) = Re (4 * (B*cnj B)\<^sup>2 / A\<^sup>2)" using Re_divide_real[of "A\<^sup>2" "4 * (B*cnj B)\<^sup>2"] \Re A > 0\ \is_real A\ by (auto simp add: power2_eq_square) moreover have "2*Re (A - D) / Re A * 2 * Re(B*cnj B) = Re (2 * (A - D) / A * 2 * B * cnj B)" using \is_real A\ \is_real D\ \A \ 0\ using Re_divide_real[of "A" "(4 * A - 4 * D) * B * cnj B"] by (simp add: field_simps) ultimately have "Re ((D - A)\<^sup>2 + 4 * B*cnj B) \ Re((A - D)\<^sup>2 + 4 * (B*cnj B)\<^sup>2 / A\<^sup>2 + 2*(A - D) / A * 2 * B*cnj B)" by (simp add: field_simps power2_eq_square) hence "Re ((D - A)\<^sup>2 + 4 * B*cnj B) \ Re(((A - D) + 2 * B*cnj B / A)\<^sup>2)" using \A \ 0\ by (subst power2_sum) (simp add: power2_eq_square field_simps) hence "dsc \ sqrt (Re(((A - D) + 2 * B*cnj B / A)\<^sup>2))" using \dsc = sqrt(Re ((D-A)\<^sup>2 + 4*(B*cnj B)))\ by simp moreover have "Re(((A - D) + 2 * B*cnj B / A)\<^sup>2) = (Re((A - D) + 2 * B*cnj B / A))\<^sup>2" using \is_real A\ \is_real D\ div_reals by (simp add: power2_eq_square) ultimately have "dsc \ \Re (A - D + 2 * B * cnj B / A)\" by simp moreover have "Re (A - D + 2 * B * cnj B / A) \ 0" proof- have *: "Re (A\<^sup>2 + B*cnj B) \ 0" using \is_real A\ by (simp add: power2_eq_square) also have "Re (A\<^sup>2 + 2*B*cnj B - A*D) \ Re (A\<^sup>2 + B*cnj B)" using \Re A * Re D \ Re (B*cnj B)\ using \is_real A\ \is_real D\ by simp finally have "Re (A\<^sup>2 + 2*B*cnj B - A*D) \ 0" by simp show ?thesis using divide_right_mono[OF \Re (A\<^sup>2 + 2*B*cnj B - A*D) \ 0\, of "Re A"] \Re A > 0\ \is_real A\ \A \ 0\ by (simp add: add_divide_distrib diff_divide_distrib) (subst Re_divide_real, auto simp add: power2_eq_square field_simps) qed ultimately have "dsc \ Re (A - D + 2 * B * cnj B / A)" by simp hence "- Re (2 * (B * cnj B) / A) \ Re ((A - D - cor dsc))" by (simp add: field_simps) hence *: "- (Re (2 * (B * cnj B)) / Re A) \ Re (A - D - cor dsc)" using \is_real A\ \A \ 0\ by (subst (asm) Re_divide_real, auto) from divide_right_mono[OF this, of "Re (2 * B * cnj B)"] have "- 1 / Re A \ Re (A - D - cor dsc) / Re (2 * B * cnj B)" using \Re A > 0\ \B \ 0\ \A \ 0\ \0 < Re (2 * (B * cnj B))\ using \(Re (2 * (B * cnj B)) / Re A) / Re (2 * B * cnj B) = 1 / Re A\ by simp from mult_right_mono[OF this, of "Re A"] show ?thesis using \is_real A\ \is_real D\ \B \ 0\ \Re A > 0\ \A \ 0\ apply (subst **) apply (subst Re_mult_real, simp) apply (subst Re_divide_real, simp, simp) apply (simp add: field_simps) done qed ultimately show ?thesis by simp next case False show ?thesis proof (cases "Re A < 0") case True hence "Re (A*a1/B) \ Re (A*a2/B)" using * ** \Re (2 * (B * cnj B)) > 0\ \B \ 0\ \is_real A\ \is_real D\ xxx using mult_right_mono_neg[of ?rhs ?lhs "Re A"] apply simp apply (subst Re_divide_real, simp, simp) apply (subst Re_divide_real, simp, simp) apply (subst Re_mult_real, simp)+ apply simp done moreover have "Re (A*a1/B) \ -1" proof- from \Re A * Re D \ Re (B*cnj B)\ have "Re (A\<^sup>2) \ Re (B*cnj B) - Re ((D - A)*A)" using \Re A < 0\ \is_real A\ \is_real D\ by (simp add: power2_eq_square field_simps) hence "1 \ Re (B*cnj B) / Re (A\<^sup>2) - Re (D - A) / Re A" using \Re A < 0\ \is_real A\ \is_real D\ using divide_right_mono[OF \Re (A\<^sup>2) \ Re (B*cnj B) - Re ((D - A)*A)\, of "Re (A\<^sup>2)"] by (simp add: power2_eq_square diff_divide_distrib) have "4 * Re(B*cnj B) \ 4 * (Re (B*cnj B))\<^sup>2 / Re (A\<^sup>2) - 2*Re (D - A) / Re A * 2 * Re(B*cnj B)" using mult_right_mono[OF \1 \ Re (B*cnj B) / Re (A\<^sup>2) - Re (D - A) / Re A\, of "4 * Re (B*cnj B)"] by (simp add: left_diff_distrib) (simp add: power2_eq_square field_simps) moreover have "A \ 0" using \Re A < 0\ by auto hence "4 * (Re (B*cnj B))\<^sup>2 / Re (A\<^sup>2) = Re (4 * (B*cnj B)\<^sup>2 / A\<^sup>2)" using Re_divide_real[of "A\<^sup>2" "4 * (B*cnj B)\<^sup>2"] \Re A < 0\ \is_real A\ by (auto simp add: power2_eq_square) moreover have "2*Re (D - A) / Re A * 2 * Re(B*cnj B) = Re (2 * (D - A) / A * 2 * B * cnj B)" using \is_real A\ \is_real D\ \A \ 0\ using Re_divide_real[of "A" "(4 * D - 4 * A) * B * cnj B"] by (simp add: field_simps) ultimately have "Re ((D - A)\<^sup>2 + 4 * B*cnj B) \ Re((D - A)\<^sup>2 + 4 * (B*cnj B)\<^sup>2 / A\<^sup>2 - 2*(D - A) / A * 2 * B*cnj B)" by (simp add: field_simps power2_eq_square) hence "Re ((D - A)\<^sup>2 + 4 * B*cnj B) \ Re(((D - A) - 2 * B*cnj B / A)\<^sup>2)" using \A \ 0\ by (subst power2_diff) (simp add: power2_eq_square field_simps) hence "dsc \ sqrt (Re(((D - A) - 2 * B*cnj B / A)\<^sup>2))" using \dsc = sqrt(Re ((D-A)\<^sup>2 + 4*(B*cnj B)))\ by simp moreover have "Re(((D - A) - 2 * B*cnj B / A)\<^sup>2) = (Re((D - A) - 2 * B*cnj B / A))\<^sup>2" using \is_real A\ \is_real D\ div_reals by (simp add: power2_eq_square) ultimately have "dsc \ \Re (D - A - 2 * B * cnj B / A)\" by simp moreover have "Re (D - A - 2 * B * cnj B / A) \ 0" proof- have "Re (A\<^sup>2 + B*cnj B) \ 0" using \is_real A\ by (simp add: power2_eq_square) also have "Re (A\<^sup>2 + 2*B*cnj B - A*D) \ Re (A\<^sup>2 + B*cnj B)" using \Re A * Re D \ Re (B*cnj B)\ using \is_real A\ \is_real D\ by simp finally have "Re (A\<^sup>2 + 2*B*cnj B - A*D) \ 0" by simp show ?thesis using divide_right_mono_neg[OF \Re (A\<^sup>2 + 2*B*cnj B - A*D) \ 0\, of "Re A"] \Re A < 0\ \is_real A\ \A \ 0\ by (simp add: add_divide_distrib diff_divide_distrib) (subst Re_divide_real, auto simp add: power2_eq_square field_simps) qed ultimately have "dsc \ Re (D - A - 2 * B * cnj B / A)" by simp hence "- Re (2 * (B * cnj B) / A) \ Re ((A - D + cor dsc))" by (simp add: field_simps) hence "- (Re (2 * (B * cnj B)) / Re A) \ Re (A - D + cor dsc)" using \is_real A\ \A \ 0\ by (subst (asm) Re_divide_real, auto) from divide_right_mono[OF this, of "Re (2 * B * cnj B)"] have "- 1 / Re A \ Re (A - D + cor dsc) / Re (2 * B * cnj B)" using \Re A < 0\ \B \ 0\ \A \ 0\ \0 < Re (2 * (B * cnj B))\ using \(Re (2 * (B * cnj B)) / Re A) / Re (2 * B * cnj B) = 1 / Re A\ by simp from mult_right_mono_neg[OF this, of "Re A"] show ?thesis using \is_real A\ \is_real D\ \B \ 0\ \Re A < 0\ \A \ 0\ apply (subst *) apply (subst Re_mult_real, simp) apply (subst Re_divide_real, simp, simp) apply (simp add: field_simps) done qed ultimately show ?thesis by simp next case False hence "A = 0" using \\ Re A > 0\ \is_real A\ using complex_eq_if_Re_eq by auto thus ?thesis by simp qed qed qed lemma chordal_circle_det_positive: fixes x y :: real assumes "x * y < 0" shows "x / (x - y) > 0" proof (cases "x > 0") case True hence "y < 0" using \x * y < 0\ by (smt mult_nonneg_nonneg) have "x - y > 0" using \x > 0\ \y < 0\ by auto thus ?thesis using \x > 0\ by (metis zero_less_divide_iff) next case False hence *: "y > 0 \ x < 0" using \x * y < 0\ using mult_nonpos_nonpos[of x y] by (cases "x=0") force+ have "x - y < 0" using * by auto thus ?thesis using * by (metis zero_less_divide_iff) qed +lemma cor_sqrt_squared: "x \ 0 \ (cor (sqrt x))\<^sup>2 = cor x" + by (simp add: power2_eq_square) + lemma chordal_circle1: assumes "is_real A" and "is_real D" and "Re (A * D) < 0" and "r = sqrt(Re ((4*A)/(A-D)))" shows "mk_circline A 0 0 D = chordal_circle \\<^sub>h r" using assms proof (transfer, transfer) fix A D r assume *: "is_real A" "is_real D" "Re (A * D) < 0" "r = sqrt (Re ((4*A)/(A-D)))" hence "A \ 0 \ D \ 0" by auto hence "(A, 0, 0, D) \ hermitean_nonzero" using eq_cnj_iff_real[of A] eq_cnj_iff_real[of D] * unfolding hermitean_def by (simp add: mat_adj_def mat_cnj_def) moreover have "(- (cor r)\<^sup>2, 0, 0, 4 - (cor r)\<^sup>2) \ hermitean_nonzero" by (simp add: hermitean_def mat_adj_def mat_cnj_def power2_eq_square) moreover have "A \ D" using \Re (A * D) < 0\ \is_real A\ \is_real D\ by auto have "Re ((4*A)/(A-D)) \ 0" proof- have "Re A / Re (A - D) \ 0" using \Re (A * D) < 0\ \is_real A\ \is_real D\ using chordal_circle_det_positive[of "Re A" "Re D"] by simp thus ?thesis using \is_real A\ \is_real D\ \A \ D\ by (subst Re_divide_real, auto) qed moreover have "- (cor (sqrt (Re (4 * A / (A - D)))))\<^sup>2 = cor (Re (4 / (D - A))) * A" - using \Re ((4*A)/(A-D)) \ 0\ \is_real A\ \is_real D\ \A \ D\ - by (subst cor_squared, subst real_sqrt_power[symmetric], simp) (simp add: Re_divide_real Re_mult_real minus_divide_right) + using \is_real A\ \is_real D\ \A \ D\ \Re ((4*A)/(A-D)) \ 0\ + by (simp add: cor_sqrt_squared field_simps) moreover - have "4 * (A - D) - 4 * A = 4 * -D" - by (simp add: field_simps) - hence "4 - 4 * A / (A - D) = -4 * D / (A - D)" - using \A \ D\ - by (smt ab_semigroup_mult_class.mult_ac(1) diff_divide_eq_iff eq_iff_diff_eq_0 mult_minus1 mult_minus1_right mult_numeral_1_right right_diff_distrib_numeral times_divide_eq_right) - hence "4 - 4 * A / (A - D) = 4 * D / (D - A)" - by (metis (hide_lams, no_types) minus_diff_eq minus_divide_left minus_divide_right minus_mult_left) + have "4 - 4 * A / (A - D) = 4 * D / (D - A)" + using\A \ D\ + by (simp add: divide_simps split: if_split_asm) (simp add: minus_mult_right) hence **: "4 - (cor (sqrt (Re (4 * A / (A - D)))))\<^sup>2 = cor (Re (4 / (D - A))) * D" using \Re ((4*A)/(A-D)) \ 0\ \is_real A\ \is_real D\ \A \ D\ - by (subst cor_squared, subst real_sqrt_power[symmetric], simp) + by (simp add: cor_sqrt_squared field_simps) ultimately show "circline_eq_cmat (mk_circline_cmat A 0 0 D) (chordal_circle_cvec_cmat \\<^sub>v r)" using * \is_real A\ \is_real D\ \A \ D\ \r = sqrt(Re ((4*A)/(A-D)))\ by (simp, rule_tac x="Re(4/(D-A))" in exI, auto, simp_all add: **) qed lemma chordal_circle2: assumes "is_real A" and "is_real D" and "Re (A * D) < 0" and "r = sqrt(Re ((4*D)/(D-A)))" shows "mk_circline A 0 0 D = chordal_circle 0\<^sub>h r" using assms proof (transfer, transfer) fix A D r assume *: "is_real A" "is_real D" "Re (A * D) < 0" "r = sqrt (Re ((4*D)/(D-A)))" hence "A \ 0 \ D \ 0" by auto hence "(A, 0, 0, D) \ {H. hermitean H \ H \ mat_zero}" using eq_cnj_iff_real[of A] eq_cnj_iff_real[of D] * unfolding hermitean_def by (simp add: mat_adj_def mat_cnj_def) moreover have "(4 - (cor r)\<^sup>2, 0, 0, - (cor r)\<^sup>2) \ {H. hermitean H \ H \ mat_zero}" by (auto simp add: hermitean_def mat_adj_def mat_cnj_def power2_eq_square) moreover have "A \ D" using \Re (A * D) < 0\ \is_real A\ \is_real D\ by auto have "Re((4*D)/(D-A)) \ 0" proof- have "Re D / Re (D - A) \ 0" using \Re (A * D) < 0\ \is_real A\ \is_real D\ using chordal_circle_det_positive[of "Re D" "Re A"] by (simp add: field_simps) thus ?thesis - using \is_real A\ \is_real D\ \A \ D\ - by (subst Re_divide_real, auto) + using \is_real A\ \is_real D\ \A \ D\ Re_divide_real by force qed - have "4 * (D - A) - 4 * D = 4 * -A" - by (simp add: field_simps) - hence "4 - 4 * D / (D - A) = -4 * A / (D - A)" - using \A \ D\ - by (smt ab_semigroup_mult_class.mult_ac(1) diff_divide_eq_iff eq_iff_diff_eq_0 mult_minus1 mult_minus1_right mult_numeral_1_right right_diff_distrib_numeral times_divide_eq_right) - hence "4 - 4 * D / (D - A) = 4 * A / (A - D)" - by (metis (hide_lams, no_types) minus_diff_eq minus_divide_left minus_divide_right minus_mult_left) + have "4 - 4 * D / (D - A) = 4 * A / (A - D)" + by (simp add: divide_simps split: if_split_asm) (simp add: \A \ D\ minus_mult_right) hence **: "4 - (cor (sqrt (Re ((4*D)/(D-A)))))\<^sup>2 = cor (Re (4 / (A - D))) * A" using \is_real A\ \is_real D\ \A \ D\ \Re (4 * D / (D - A)) \ 0\ - by (subst cor_squared, subst real_sqrt_power[symmetric], simp) - + by (simp add: cor_sqrt_squared field_simps) moreover have "- (cor (sqrt (Re ((4*D)/(D-A)))))\<^sup>2 = cor (Re (4 / (A - D))) * D" - using \is_real A\ \is_real D\ \A \ D\ \Re ((4*D)/(D-A)) \ 0\ - by (subst cor_squared, subst real_sqrt_power[symmetric], simp) (simp add: Re_divide_real minus_divide_right) - + using \is_real A\ \is_real D\ \A \ D\ \Re (4 * D / (D - A)) \ 0\ + by (simp add: cor_sqrt_squared field_simps) ultimately show "circline_eq_cmat (mk_circline_cmat A 0 0 D) (chordal_circle_cvec_cmat 0\<^sub>v r)" using \is_real A\ \is_real D\ \A \ 0 \ D \ 0\ \r = sqrt (Re ((4*D)/(D-A)))\ using * by (simp, rule_tac x="Re (4/(A-D))" in exI, auto, simp_all add: **) qed lemma chordal_circle': assumes "B \ 0" and "(A, B, C, D) \ hermitean_nonzero" and "Re (mat_det (A, B, C, D)) \ 0" and "C * a\<^sup>2 + (D - A) * a - B = 0" and "r = sqrt((4 - Re((-4 * a/B) * A)) / (1 + Re (a*cnj a)))" shows "mk_circline A B C D = chordal_circle (of_complex a) r" using assms proof (transfer, transfer) fix A B C D a :: complex and r :: real let ?k = "(-4) * a / B" assume *: "(A, B, C, D) \ {H. hermitean H \ H \ mat_zero}" and **: "B \ 0" "C * a\<^sup>2 + (D - A) * a - B = 0" and rr: "r = sqrt ((4 - Re (?k * A)) / (1 + Re (a * cnj a)))" and det: "Re (mat_det (A, B, C, D)) \ 0" have "is_real A" "is_real D" "C = cnj B" using * hermitean_elems by auto from ** have a12: "let dsc = sqrt(Re ((D-A)\<^sup>2 + 4 * (B*cnj B))) in a = (A - D + cor dsc) / (2 * C) \ a = (A - D - cor dsc) / (2 * C)" proof- have "Re ((D-A)\<^sup>2 + 4 * (B*cnj B)) \ 0" using \is_real A\ \is_real D\ by (subst complex_mult_cnj_cmod) (simp add: power2_eq_square) hence "ccsqrt ((D - A)\<^sup>2 - 4 * C * - B) = cor (sqrt (Re ((D - A)\<^sup>2 + 4 * (B * cnj B))))" using csqrt_real[of "((D - A)\<^sup>2 + 4 * (B * cnj B))"] \is_real A\ \is_real D\ \C = cnj B\ by (auto simp add: power2_eq_square field_simps) thus ?thesis using complex_quadratic_equation_two_roots[of C a "D - A" "-B"] using \C * a\<^sup>2 + (D - A) * a - B = 0\ \B \ 0\ \C = cnj B\ by (simp add: Let_def) qed have "is_real ?k" using a12 \C = cnj B\ \is_real A\ \is_real D\ by (auto simp add: Let_def) have "a \ 0" using ** by auto hence "Re ?k \ 0" using \is_real (-4*a / B)\ \B \ 0\ by (metis complex.expand divide_eq_0_iff divisors_zero zero_complex.simps(1) zero_complex.simps(2) zero_neq_neg_numeral) moreover have "(-4) * a = cor (Re ?k) * B" using complex_of_real_Re[OF \is_real (-4*a/B)\] \B \ 0\ by simp moreover have "is_real (a/B)" using \is_real ?k\ is_real_mult_real[of "-4" "a / B"] by simp hence "is_real (B * cnj a)" using * \C = cnj B\ by (metis (no_types, lifting) Im_complex_div_eq_0 complex_cnj_divide eq_cnj_iff_real hermitean_elems(3) mem_Collect_eq mult.commute) hence "B * cnj a = cnj B * a" using eq_cnj_iff_real[of "B * cnj a"] by simp hence "-4 * cnj a = cor (Re ?k) * C" using \C = cnj B\ using complex_of_real_Re[OF \is_real ?k\] \B \ 0\ by (simp, simp add: field_simps) moreover have "1 + a * cnj a \ 0" - by (subst complex_mult_cnj_cmod) (smt cor_add of_real_0 of_real_1 of_real_eq_iff realpow_square_minus_le) + by (simp add: complex_mult_cnj_cmod) have "r\<^sup>2 = (4 - Re (?k * A)) / (1 + Re (a * cnj a))" proof- have "Re (a / B * A) \ -1" using a12 chordal_circle_radius_positive[of A B C D] * \B \ 0\ det by (auto simp add: Let_def field_simps) from mult_right_mono_neg[OF this, of "-4"] have "4 - Re (?k * A) \ 0" using Re_mult_real[of "-4" "a / B * A"] by (simp add: field_simps) moreover have "1 + Re (a * cnj a) > 0" using \a \ 0\ complex_mult_cnj complex_neq_0 by auto ultimately have "(4 - Re (?k * A)) / (1 + Re (a * cnj a)) \ 0" by (metis divide_nonneg_pos) thus ?thesis using rr by simp qed hence "r\<^sup>2 = Re ((4 - ?k * A) / (1 + a * cnj a))" using \is_real ?k\ \is_real A\ \1 + a * cnj a \ 0\ by (subst Re_divide_real, auto) hence "(cor r)\<^sup>2 = (4 - ?k * A) / (1 + a * cnj a)" - using \is_real ?k\ \is_real A\ - using mult_reals[of ?k A] + using \is_real ?k\ \is_real A\ mult_reals[of ?k A] by (simp add: cor_squared) hence "4 - (cor r)\<^sup>2 * (a * cnj a + 1) = cor (Re ?k) * A" using complex_of_real_Re[OF \is_real (-4*a/B)\] using \1 + a * cnj a \ 0\ by (simp add: field_simps) moreover have "?k = cnj ?k" using \is_real ?k\ using eq_cnj_iff_real[of "-4*a/B"] by simp have "?k\<^sup>2 = cor ((cmod ?k)\<^sup>2)" using cor_cmod_real[OF \is_real ?k\] - unfolding power2_eq_square - by (subst cor_mult) (metis minus_mult_minus) + unfolding power2_eq_square by force hence "?k\<^sup>2 = ?k * cnj ?k" using complex_mult_cnj_cmod[of ?k] by simp hence ***: "a * cnj a = (cor ((Re ?k)\<^sup>2) * B * C) / 16" using complex_of_real_Re[OF \is_real (-4*a/B)\] \C = cnj B\ \is_real (-4*a/B)\ \B \ 0\ by simp from ** have "cor ((Re ?k)\<^sup>2) * B * C - 4 * cor (Re ?k) * (D-A) - 16 = 0" using complex_of_real_Re[OF \is_real ?k\] by (simp add: power2_eq_square, simp add: field_simps, algebra) hence "?k * (D-A) = 4 * (cor ((Re ?k)\<^sup>2) * B * C / 16 - 1)" by (subst (asm) complex_of_real_Re[OF \is_real ?k\]) algebra hence "?k * (D-A) = 4 * (a*cnj a - 1)" by (subst (asm) ***[symmetric]) simp hence "4 * a * cnj a - (cor r)\<^sup>2 * (a * cnj a + 1) = cor (Re ?k) * D" using \4 - (cor r)\<^sup>2 * (a * cnj a + 1) = cor (Re ?k) * A\ using complex_of_real_Re[OF \is_real (-4*a/B)\] by simp algebra ultimately show "circline_eq_cmat (mk_circline_cmat A B C D) (chordal_circle_cvec_cmat (of_complex_cvec a) r)" using * \a \ 0\ by (simp, rule_tac x="Re (-4*a / B)" in exI, simp) qed end diff --git a/thys/Complex_Geometry/Circlines.thy b/thys/Complex_Geometry/Circlines.thy --- a/thys/Complex_Geometry/Circlines.thy +++ b/thys/Complex_Geometry/Circlines.thy @@ -1,2048 +1,2048 @@ (* ---------------------------------------------------------------------------- *) section \Circlines\ (* ---------------------------------------------------------------------------- *) theory Circlines imports More_Set Moebius Hermitean_Matrices Elementary_Complex_Geometry begin (* ----------------------------------------------------------------- *) subsection \Definition of circlines\ (* ----------------------------------------------------------------- *) text \In our formalization we follow the approach described by Schwerdtfeger \cite{schwerdtfeger} and represent circlines by Hermitean, non-zero $2\times 2$ matrices. In the original formulation, a matrix $\left(\begin{array}{cc}A & B\\C & D\end{array}\right)$ corresponds to the equation $A\cdot z\cdot \overline{z} + B\cdot \overline{z} + C\cdot z + D = 0$, where $C = \overline{B}$ and $A$ and $D$ are real (as the matrix is Hermitean).\ abbreviation hermitean_nonzero where "hermitean_nonzero \ {H. hermitean H \ H \ mat_zero}" typedef circline_mat = hermitean_nonzero by (rule_tac x="eye" in exI) (auto simp add: hermitean_def mat_adj_def mat_cnj_def) setup_lifting type_definition_circline_mat definition circline_eq_cmat :: "complex_mat \ complex_mat \ bool" where [simp]: "circline_eq_cmat A B \ (\ k::real. k \ 0 \ B = cor k *\<^sub>s\<^sub>m A)" lemma symp_circline_eq_cmat: "symp circline_eq_cmat" unfolding symp_def proof ((rule allI)+, rule impI) fix x y assume "circline_eq_cmat x y" then obtain k where "k \ 0 \ y = cor k *\<^sub>s\<^sub>m x" by auto hence "1 / k \ 0 \ x = cor (1 / k) *\<^sub>s\<^sub>m y" by auto thus "circline_eq_cmat y x" unfolding circline_eq_cmat_def by blast qed text\Hermitean non-zero matrices are equivalent only to such matrices\ lemma circline_eq_cmat_hermitean_nonzero: assumes "hermitean H \ H \ mat_zero" "circline_eq_cmat H H'" shows "hermitean H' \ H' \ mat_zero" using assms by (metis circline_eq_cmat_def hermitean_mult_real nonzero_mult_real of_real_eq_0_iff) lift_definition circline_eq_clmat :: "circline_mat \ circline_mat \ bool" is circline_eq_cmat done lemma circline_eq_clmat_refl [simp]: "circline_eq_clmat H H" by transfer (simp, rule_tac x="1" in exI, simp) quotient_type circline = circline_mat / circline_eq_clmat proof (rule equivpI) show "reflp circline_eq_clmat" unfolding reflp_def by transfer (auto, rule_tac x="1" in exI, simp) next show "symp circline_eq_clmat" unfolding symp_def by transfer (auto, (rule_tac x="1/k" in exI, simp)+) next show "transp circline_eq_clmat" unfolding transp_def by transfer (simp, safe, (rule_tac x="ka*k" in exI, simp)+) qed text \Circline with specified matrix\ text \An auxiliary constructor @{term mk_circline} returns a circline (an equivalence class) for given four complex numbers $A$, $B$, $C$ and $D$ (provided that they form a Hermitean, non-zero matrix).\ definition mk_circline_cmat :: "complex \ complex \ complex \ complex \ complex_mat" where [simp]: "mk_circline_cmat A B C D = (let M = (A, B, C, D) in if M \ hermitean_nonzero then M else eye)" lift_definition mk_circline_clmat :: "complex \ complex \ complex \ complex \ circline_mat" is mk_circline_cmat by (auto simp add: Let_def hermitean_def mat_adj_def mat_cnj_def) lift_definition mk_circline :: "complex \ complex \ complex \ complex \ circline" is mk_circline_clmat done lemma ex_mk_circline: shows "\ A B C D. H = mk_circline A B C D \ hermitean (A, B, C, D) \ (A, B, C, D) \ mat_zero" proof (transfer, transfer) fix H assume *: "hermitean H \ H \ mat_zero" obtain A B C D where "H = (A, B, C, D)" by (cases " H", auto) hence "circline_eq_cmat H (mk_circline_cmat A B C D) \ hermitean (A, B, C, D) \ (A, B, C, D) \ mat_zero" using * by auto thus "\ A B C D. circline_eq_cmat H (mk_circline_cmat A B C D) \ hermitean (A, B, C, D) \ (A, B, C, D) \ mat_zero" by blast qed (* ----------------------------------------------------------------- *) subsection \Circline type\ (* ----------------------------------------------------------------- *) definition circline_type_cmat :: "complex_mat \ real" where [simp]: "circline_type_cmat H = sgn (Re (mat_det H))" lift_definition circline_type_clmat :: "circline_mat \ real" is circline_type_cmat done lift_definition circline_type :: "circline \ real" is circline_type_clmat by transfer (simp, erule exE, simp add: sgn_mult) lemma circline_type: "circline_type H = -1 \ circline_type H = 0 \ circline_type H = 1" by (transfer, transfer, simp add: sgn_if) lemma circline_type_mk_circline [simp]: assumes "(A, B, C, D) \ hermitean_nonzero" shows "circline_type (mk_circline A B C D) = sgn (Re (A*D - B*C))" using assms by (transfer, transfer, simp) (* ----------------------------------------------------------------- *) subsection \Points on the circline\ (* ----------------------------------------------------------------- *) text \Each circline determines a corresponding set of points. Again, a description given in homogeneous coordinates is a bit better than the original description defined only for ordinary complex numbers. The point with homogeneous coordinates $(z_1, z_2)$ will belong to the set of circline points iff $A \cdot z_1\cdot \overline{z_1} + B\cdot \overline{z_1} \cdot z_2 + C\cdot z_1 \cdot\overline{z_2} + D\cdot z_2 \cdot \overline{z_2} = 0$. Note that this is a quadratic form determined by a vector of homogeneous coordinates and the Hermitean matrix.\ definition on_circline_cmat_cvec :: "complex_mat \ complex_vec \ bool" where [simp]: "on_circline_cmat_cvec H z \ quad_form z H = 0" lift_definition on_circline_clmat_hcoords :: "circline_mat \ complex_homo_coords \ bool" is on_circline_cmat_cvec done lift_definition on_circline :: "circline \ complex_homo \ bool" is on_circline_clmat_hcoords by transfer (simp del: quad_form_def, (erule exE)+, simp del: quad_form_def add: quad_form_scale_m quad_form_scale_v) definition circline_set :: "circline \ complex_homo set" where "circline_set H = {z. on_circline H z}" lemma circline_set_I [simp]: assumes "on_circline H z" shows "z \ circline_set H" using assms unfolding circline_set_def by auto abbreviation circline_equation where "circline_equation A B C D z1 z2 \ A*z1*cnj z1 + B*z2*cnj z1 + C*cnj z2*z1 + D*z2*cnj z2 = 0" lemma on_circline_cmat_cvec_circline_equation: "on_circline_cmat_cvec (A, B, C, D) (z1, z2) \ circline_equation A B C D z1 z2" by (simp add: vec_cnj_def field_simps) lemma circline_equation: assumes "H = mk_circline A B C D" and "(A, B, C, D) \ hermitean_nonzero" shows "of_complex z \ circline_set H \ circline_equation A B C D z 1" using assms unfolding circline_set_def by simp (transfer, transfer, simp add: vec_cnj_def field_simps) text \Circlines trough 0 and inf.\ text \The circline represents a line when $A=0$ or a circle, otherwise.\ definition circline_A0_cmat :: "complex_mat \ bool" where [simp]: "circline_A0_cmat H \ (let (A, B, C, D) = H in A = 0)" lift_definition circline_A0_clmat :: "circline_mat \ bool" is circline_A0_cmat done lift_definition circline_A0 :: "circline \ bool" is circline_A0_clmat by transfer auto abbreviation is_line where "is_line H \ circline_A0 H" abbreviation is_circle where "is_circle H \ \ circline_A0 H" definition circline_D0_cmat :: "complex_mat \ bool" where [simp]: "circline_D0_cmat H \ (let (A, B, C, D) = H in D = 0)" lift_definition circline_D0_clmat :: "circline_mat \ bool" is circline_D0_cmat done lift_definition circline_D0 :: "circline \ bool" is circline_D0_clmat by transfer auto lemma inf_on_circline: "on_circline H \\<^sub>h \ circline_A0 H" by (transfer, transfer, auto simp add: vec_cnj_def) lemma inf_in_circline_set: "\\<^sub>h \ circline_set H \ is_line H" using inf_on_circline unfolding circline_set_def by simp lemma zero_on_circline: "on_circline H 0\<^sub>h \ circline_D0 H" by (transfer, transfer, auto simp add: vec_cnj_def) lemma zero_in_circline_set: "0\<^sub>h \ circline_set H \ circline_D0 H" using zero_on_circline unfolding circline_set_def by simp (* ----------------------------------------------------------------- *) subsection \Connection with circles and lines in the classic complex plane\ (* ----------------------------------------------------------------- *) text \Every Euclidean circle and Euclidean line can be represented by a circline.\ lemma classic_circline: assumes "H = mk_circline A B C D" and "hermitean (A, B, C, D) \ (A, B, C, D) \ mat_zero" shows "circline_set H - {\\<^sub>h} = of_complex ` circline (Re A) B (Re D)" using assms unfolding circline_set_def proof (safe) fix z assume "hermitean (A, B, C, D)" "(A, B, C, D) \ mat_zero" "z \ circline (Re A) B (Re D)" thus "on_circline (mk_circline A B C D) (of_complex z)" using hermitean_elems[of A B C D] by (transfer, transfer) (auto simp add: circline_def vec_cnj_def field_simps) next fix z assume "of_complex z = \\<^sub>h" thus False by simp next fix z assume "hermitean (A, B, C, D)" "(A, B, C, D) \ mat_zero" "on_circline (mk_circline A B C D) z" "z \ of_complex ` circline (Re A) B (Re D)" moreover have "z \ \\<^sub>h \ z \ of_complex ` circline (Re A) B (Re D)" proof assume "z \ \\<^sub>h" show "z \ of_complex ` circline (Re A) B (Re D)" proof show "z = of_complex (to_complex z)" using \z \ \\<^sub>h\ by simp next show "to_complex z \ circline (Re A) B (Re D)" using \on_circline (mk_circline A B C D) z\ \z \ \\<^sub>h\ using \hermitean (A, B, C, D)\ \(A, B, C, D) \ mat_zero\ proof (transfer, transfer) fix A B C D and z :: complex_vec obtain z1 z2 where zz: "z = (z1, z2)" by (cases z, auto) assume *: "z \ vec_zero" "\ z \\<^sub>v \\<^sub>v" "on_circline_cmat_cvec (mk_circline_cmat A B C D) z" "hermitean (A, B, C, D)" "(A, B, C, D) \ mat_zero" have "z2 \ 0" using \z \ vec_zero\ \\ z \\<^sub>v \\<^sub>v\ using inf_cvec_z2_zero_iff zz by blast thus "to_complex_cvec z \ circline (Re A) B (Re D)" using * zz using hermitean_elems[of A B C D] by (simp add: vec_cnj_def circline_def field_simps) qed qed qed ultimately show "z = \\<^sub>h" by simp qed text \The matrix of the circline representing circle determined with center and radius.\ definition mk_circle_cmat :: "complex \ real \ complex_mat" where [simp]: "mk_circle_cmat a r = (1, -a, -cnj a, a*cnj a - cor r*cor r)" lift_definition mk_circle_clmat :: "complex \ real \ circline_mat" is mk_circle_cmat by (simp add: hermitean_def mat_adj_def mat_cnj_def) lift_definition mk_circle :: "complex \ real \ circline" is mk_circle_clmat done lemma is_circle_mk_circle: "is_circle (mk_circle a r)" by (transfer, transfer, simp) lemma circline_set_mk_circle [simp]: assumes "r \ 0" shows "circline_set (mk_circle a r) = of_complex ` circle a r" proof- let ?A = "1" and ?B = "-a" and ?C = "-cnj a" and ?D = "a*cnj a - cor r*cor r" have *: "(?A, ?B, ?C, ?D) \ {H. hermitean H \ H \ mat_zero}" by (simp add: hermitean_def mat_adj_def mat_cnj_def) have "mk_circle a r = mk_circline ?A ?B ?C ?D" using * by (transfer, transfer, simp) hence "circline_set (mk_circle a r) - {\\<^sub>h} = of_complex ` circline ?A ?B (Re ?D)" using classic_circline[of "mk_circle a r" ?A ?B ?C ?D] * by simp moreover have "circline ?A ?B (Re ?D) = circle a r" by (rule circline_circle[of ?A "Re ?D" "?B" "circline ?A ?B (Re ?D)" "a" "r*r" r], simp_all add: cmod_square \r \ 0\) moreover have "\\<^sub>h \ circline_set (mk_circle a r)" using inf_in_circline_set[of "mk_circle a r"] is_circle_mk_circle[of a r] by auto ultimately show ?thesis unfolding circle_def by simp qed text \The matrix of the circline representing line determined with two (not equal) complex points.\ definition mk_line_cmat :: "complex \ complex \ complex_mat" where [simp]: "mk_line_cmat z1 z2 = (if z1 \ z2 then let B = \ * (z2 - z1) in (0, B, cnj B, -cnj_mix B z1) else eye)" lift_definition mk_line_clmat :: "complex \ complex \ circline_mat" is mk_line_cmat by (auto simp add: Let_def hermitean_def mat_adj_def mat_cnj_def split: if_split_asm) lift_definition mk_line :: "complex \ complex \ circline" is mk_line_clmat done lemma circline_set_mk_line [simp]: assumes "z1 \ z2" shows "circline_set (mk_line z1 z2) - {\\<^sub>h} = of_complex ` line z1 z2" proof- let ?A = "0" and ?B = "\*(z2 - z1)" let ?C = "cnj ?B" and ?D = "-cnj_mix ?B z1" have *: "(?A, ?B, ?C, ?D) \ {H. hermitean H \ H \ mat_zero}" using assms by (simp add: hermitean_def mat_adj_def mat_cnj_def) have "mk_line z1 z2 = mk_circline ?A ?B ?C ?D" using * assms by (transfer, transfer, auto simp add: Let_def) hence "circline_set (mk_line z1 z2) - {\\<^sub>h} = of_complex ` circline ?A ?B (Re ?D)" using classic_circline[of "mk_line z1 z2" ?A ?B ?C ?D] * by simp moreover have "circline ?A ?B (Re ?D) = line z1 z2" using \z1 \ z2\ using circline_line' by simp ultimately show ?thesis by simp qed text \The set of points determined by a circline is always either an Euclidean circle or an Euclidean line. \ text \Euclidean circle is determined by its center and radius.\ type_synonym euclidean_circle = "complex \ real" definition euclidean_circle_cmat :: "complex_mat \ euclidean_circle" where [simp]: "euclidean_circle_cmat H = (let (A, B, C, D) = H in (-B/A, sqrt(Re ((B*C - A*D)/(A*A)))))" lift_definition euclidean_circle_clmat :: "circline_mat \ euclidean_circle" is euclidean_circle_cmat done lift_definition euclidean_circle :: "circline \ euclidean_circle" is euclidean_circle_clmat proof transfer fix H1 H2 assume hh: "hermitean H1 \ H1 \ mat_zero" "hermitean H2 \ H2 \ mat_zero" obtain A1 B1 C1 D1 where HH1: "H1 = (A1, B1, C1, D1)" by (cases "H1") auto obtain A2 B2 C2 D2 where HH2: "H2 = (A2, B2, C2, D2)" by (cases "H2") auto assume "circline_eq_cmat H1 H2" then obtain k where "k \ 0" and *: "A2 = cor k * A1" "B2 = cor k * B1" "C2 = cor k * C1" "D2 = cor k * D1" using HH1 HH2 by auto have "(cor k * B1 * (cor k * C1) - cor k * A1 * (cor k * D1)) = (cor k)\<^sup>2 * (B1*C1 - A1*D1)" "(cor k * A1 * (cor k * A1)) = (cor k)\<^sup>2 * (A1*A1)" by (auto simp add: field_simps power2_eq_square) hence "(cor k * B1 * (cor k * C1) - cor k * A1 * (cor k * D1)) / (cor k * A1 * (cor k * A1)) = (B1*C1 - A1*D1) / (A1*A1)" using \k \ 0\ by (simp add: power2_eq_square) thus "euclidean_circle_cmat H1 = euclidean_circle_cmat H2" using HH1 HH2 * hh by auto qed lemma classic_circle: assumes "is_circle H" and "(a, r) = euclidean_circle H" and "circline_type H \ 0" shows "circline_set H = of_complex ` circle a r" proof- obtain A B C D where *: "H = mk_circline A B C D" "hermitean (A, B, C, D)" "(A, B, C, D) \ mat_zero" using ex_mk_circline[of H] by auto have "is_real A" "is_real D" "C = cnj B" using * hermitean_elems by auto have "Re (A*D - B*C) \ 0" using \circline_type H \ 0\ * by simp hence **: "Re A * Re D \ (cmod B)\<^sup>2" using \is_real A\ \is_real D\ \C = cnj B\ by (simp add: cmod_square) have "A \ 0" using \is_circle H\ * \is_real A\ by simp (transfer, transfer, simp) hence "Re A \ 0" using \is_real A\ by (metis complex_surj zero_complex.code) have ***: "\\<^sub>h \ circline_set H" using * inf_in_circline_set[of H] \is_circle H\ by simp let ?a = "-B/A" let ?r2 = "((cmod B)\<^sup>2 - Re A * Re D) / (Re A)\<^sup>2" let ?r = "sqrt ?r2" have "?a = a \ ?r = r" using \(a, r) = euclidean_circle H\ using * \is_real A\ \is_real D\ \C = cnj B\ \A \ 0\ apply simp apply transfer apply transfer apply simp apply (subst Re_divide_real) apply (simp_all add: cmod_square, simp add: power2_eq_square) done show ?thesis using * ** *** \Re A \ 0\ \is_real A\ \C = cnj B\ \?a = a \ ?r = r\ using classic_circline[of H A B C D] assms circline_circle[of "Re A" "Re D" B "circline (Re A) B (Re D)" ?a ?r2 ?r] by (simp add: circle_def) qed text \Euclidean line is represented by two points.\ type_synonym euclidean_line = "complex \ complex" definition euclidean_line_cmat :: "complex_mat \ euclidean_line" where [simp]: "euclidean_line_cmat H = (let (A, B, C, D) = H; z1 = -(D*B)/(2*B*C); z2 = z1 + \ * sgn (if arg B > 0 then -B else B) in (z1, z2))" lift_definition euclidean_line_clmat :: "circline_mat \ euclidean_line" is euclidean_line_cmat done lift_definition euclidean_line :: "circline \ complex \ complex" is euclidean_line_clmat proof transfer fix H1 H2 assume hh: "hermitean H1 \ H1 \ mat_zero" "hermitean H2 \ H2 \ mat_zero" obtain A1 B1 C1 D1 where HH1: "H1 = (A1, B1, C1, D1)" by (cases "H1") auto obtain A2 B2 C2 D2 where HH2: "H2 = (A2, B2, C2, D2)" by (cases "H2") auto assume "circline_eq_cmat H1 H2" then obtain k where "k \ 0" and *: "A2 = cor k * A1" "B2 = cor k * B1" "C2 = cor k * C1" "D2 = cor k * D1" using HH1 HH2 by auto have 1: "B1 \ 0 \ 0 < arg B1 \ \ 0 < arg (- B1)" using canon_ang_plus_pi1[of "arg B1"] arg_bounded[of B1] by (auto simp add: arg_uminus) have 2: "B1 \ 0 \ \ 0 < arg B1 \ 0 < arg (- B1)" using canon_ang_plus_pi2[of "arg B1"] arg_bounded[of B1] by (auto simp add: arg_uminus) show "euclidean_line_cmat H1 = euclidean_line_cmat H2" using HH1 HH2 * \k \ 0\ - by (cases "k > 0") (auto simp add: Let_def, simp_all add: sgn_eq 1 2) + by (cases "k > 0") (auto simp add: Let_def, simp_all add: norm_mult sgn_eq 1 2) qed lemma classic_line: assumes "is_line H" and "circline_type H < 0" and "(z1, z2) = euclidean_line H" shows "circline_set H - {\\<^sub>h} = of_complex ` line z1 z2" proof- obtain A B C D where *: "H = mk_circline A B C D" "hermitean (A, B, C, D)" "(A, B, C, D) \ mat_zero" using ex_mk_circline[of H] by auto have "is_real A" "is_real D" "C = cnj B" using * hermitean_elems by auto have "Re A = 0" using \is_line H\ * \is_real A\ \is_real D\ \C = cnj B\ by simp (transfer, transfer, simp) have "B \ 0" using \Re A = 0\ \is_real A\ \is_real D\ \C = cnj B\ * \circline_type H < 0\ using circline_type_mk_circline[of A B C D] by auto let ?z1 = "- cor (Re D) * B / (2 * B * cnj B)" let ?z2 = "?z1 + \ * sgn (if 0 < arg B then - B else B)" have "z1 = ?z1 \ z2 = ?z2" using \(z1, z2) = euclidean_line H\ * \is_real A\ \is_real D\ \C = cnj B\ by simp (transfer, transfer, simp add: Let_def) thus ?thesis using * using classic_circline[of H A B C D] circline_line[of "Re A" B "circline (Re A) B (Re D)" "Re D" ?z1 ?z2] \Re A = 0\ \B \ 0\ by simp qed (* ----------------------------------------------------------------- *) subsection \Some special circlines\ (* ----------------------------------------------------------------- *) (* ---------------------------------------------------------------------------- *) subsubsection \Unit circle\ (* ---------------------------------------------------------------------------- *) definition unit_circle_cmat :: complex_mat where [simp]: "unit_circle_cmat = (1, 0, 0, -1)" lift_definition unit_circle_clmat :: circline_mat is unit_circle_cmat by (simp add: hermitean_def mat_adj_def mat_cnj_def) lift_definition unit_circle :: circline is unit_circle_clmat done lemma on_circline_cmat_cvec_unit: shows "on_circline_cmat_cvec unit_circle_cmat (z1, z2) \ z1 * cnj z1 = z2 * cnj z2" by (simp add: vec_cnj_def field_simps) lemma one_on_unit_circle [simp]: "on_circline unit_circle 1\<^sub>h" and ii_on_unit_circle [simp]: "on_circline unit_circle ii\<^sub>h" and not_zero_on_unit_circle [simp]: "\ on_circline unit_circle 0\<^sub>h" by (transfer, transfer, simp add: vec_cnj_def)+ lemma one_in_unit_circle_set [simp]: "1\<^sub>h \ circline_set unit_circle" and ii_in_unit_circle_set [simp]: "ii\<^sub>h \ circline_set unit_circle" and zero_in_unit_circle_set [simp]: "0\<^sub>h \ circline_set unit_circle" unfolding circline_set_def by simp_all lemma is_circle_unit_circle [simp]: shows "is_circle unit_circle" by (transfer, transfer, simp) lemma not_inf_on_unit_circle' [simp]: shows "\ on_circline unit_circle \\<^sub>h" using is_circle_unit_circle inf_on_circline by blast lemma not_inf_on_unit_circle'' [simp]: shows "\\<^sub>h \ circline_set unit_circle" by (simp add: inf_in_circline_set) lemma euclidean_circle_unit_circle [simp]: shows "euclidean_circle unit_circle = (0, 1)" by (transfer, transfer, simp) lemma circline_type_unit_circle [simp]: shows "circline_type unit_circle = -1" by (transfer, transfer, simp) lemma on_circline_unit_circle [simp]: shows "on_circline unit_circle (of_complex z) \ cmod z = 1" by (transfer, transfer, simp add: vec_cnj_def mult.commute) lemma circline_set_unit_circle [simp]: shows "circline_set unit_circle = of_complex ` {z. cmod z = 1}" proof- show ?thesis proof safe fix x assume "x \ circline_set unit_circle" then obtain x' where "x = of_complex x'" using inf_or_of_complex[of x] by auto thus "x \ of_complex ` {z. cmod z = 1}" using \x \ circline_set unit_circle\ unfolding circline_set_def by auto next fix x assume "cmod x = 1" thus "of_complex x \ circline_set unit_circle" unfolding circline_set_def by auto qed qed lemma circline_set_unit_circle_I [simp]: assumes "cmod z = 1" shows "of_complex z \ circline_set unit_circle" using assms unfolding circline_set_unit_circle by simp lemma inversion_unit_circle [simp]: assumes "on_circline unit_circle x" shows "inversion x = x" proof- obtain x' where "x = of_complex x'" "x' \ 0" using inf_or_of_complex[of x] using assms by force moreover hence "x' * cnj x' = 1" using assms using circline_set_unit_circle unfolding circline_set_def by auto hence "1 / cnj x' = x'" using \x' \ 0\ by (simp add: field_simps) ultimately show ?thesis using assms unfolding inversion_def by simp qed lemma inversion_id_iff_on_unit_circle: shows "inversion a = a \ on_circline unit_circle a" using inversion_id_iff[of a] inf_or_of_complex[of a] by auto lemma on_unit_circle_conjugate [simp]: shows "on_circline unit_circle (conjugate z) \ on_circline unit_circle z" by (transfer, transfer, auto simp add: vec_cnj_def field_simps) lemma conjugate_unit_circle_set [simp]: shows "conjugate ` (circline_set unit_circle) = circline_set unit_circle" unfolding circline_set_def by (auto simp add: image_iff, rule_tac x="conjugate x" in exI, simp) (* ---------------------------------------------------------------------------- *) subsubsection \x-axis\ (* ---------------------------------------------------------------------------- *) definition x_axis_cmat :: complex_mat where [simp]: "x_axis_cmat = (0, \, -\, 0)" lift_definition x_axis_clmat :: circline_mat is x_axis_cmat by (simp add: hermitean_def mat_adj_def mat_cnj_def) lift_definition x_axis :: circline is x_axis_clmat done lemma special_points_on_x_axis' [simp]: shows "on_circline x_axis 0\<^sub>h" and "on_circline x_axis 1\<^sub>h" and "on_circline x_axis \\<^sub>h" by (transfer, transfer, simp add: vec_cnj_def)+ lemma special_points_on_x_axis'' [simp]: shows "0\<^sub>h \ circline_set x_axis" and "1\<^sub>h \ circline_set x_axis" and "\\<^sub>h \ circline_set x_axis" unfolding circline_set_def by auto lemma is_line_x_axis [simp]: shows "is_line x_axis" by (transfer, transfer, simp) lemma circline_type_x_axis [simp]: shows "circline_type x_axis = -1" by (transfer, transfer, simp) lemma on_circline_x_axis: shows "on_circline x_axis z \ (\ c. is_real c \ z = of_complex c) \ z = \\<^sub>h" proof safe fix z c assume "is_real c" thus "on_circline x_axis (of_complex c)" proof (transfer, transfer) fix c assume "is_real c" thus "on_circline_cmat_cvec x_axis_cmat (of_complex_cvec c)" using eq_cnj_iff_real[of c] by (simp add: vec_cnj_def) qed next fix z assume "on_circline x_axis z" "z \ \\<^sub>h" thus "\c. is_real c \ z = of_complex c" proof (transfer, transfer, safe) fix a b assume "(a, b) \ vec_zero" "on_circline_cmat_cvec x_axis_cmat (a, b)" "\ (a, b) \\<^sub>v \\<^sub>v" hence "b \ 0" "cnj a * b = cnj b * a" using inf_cvec_z2_zero_iff by (auto simp add: vec_cnj_def) thus "\c. is_real c \ (a, b) \\<^sub>v of_complex_cvec c" apply (rule_tac x="a/b" in exI) apply (auto simp add: is_real_div field_simps) apply (rule_tac x="1/b" in exI, simp) done qed next show "on_circline x_axis \\<^sub>h" by auto qed lemma on_circline_x_axis_I [simp]: assumes "is_real z" shows "on_circline x_axis (of_complex z)" using assms unfolding on_circline_x_axis by auto lemma circline_set_x_axis: shows "circline_set x_axis = of_complex ` {x. is_real x} \ {\\<^sub>h}" using on_circline_x_axis unfolding circline_set_def by auto lemma circline_set_x_axis_I: assumes "is_real z" shows "of_complex z \ circline_set x_axis" using assms unfolding circline_set_x_axis by auto lemma circline_equation_x_axis: shows "of_complex z \ circline_set x_axis \ z = cnj z" unfolding circline_set_x_axis proof auto fix x assume "of_complex z = of_complex x" "is_real x" hence "z = x" using of_complex_inj[of z x] by simp thus "z = cnj z" using eq_cnj_iff_real[of z] \is_real x\ by auto next assume "z = cnj z" thus "of_complex z \ of_complex ` {x. is_real x} " using eq_cnj_iff_real[of z] by auto qed text \Positive and negative part of x-axis\ definition positive_x_axis where "positive_x_axis = {z. z \ circline_set x_axis \ z \ \\<^sub>h \ Re (to_complex z) > 0}" definition negative_x_axis where "negative_x_axis = {z. z \ circline_set x_axis \ z \ \\<^sub>h \ Re (to_complex z) < 0}" lemma circline_set_positive_x_axis_I [simp]: assumes "is_real z" and "Re z > 0" shows "of_complex z \ positive_x_axis" using assms unfolding positive_x_axis_def by simp lemma circline_set_negative_x_axis_I [simp]: assumes "is_real z" and "Re z < 0" shows "of_complex z \ negative_x_axis" using assms unfolding negative_x_axis_def by simp (* ---------------------------------------------------------------------------- *) subsubsection \y-axis\ (* ---------------------------------------------------------------------------- *) definition y_axis_cmat :: complex_mat where [simp]: "y_axis_cmat = (0, 1, 1, 0)" lift_definition y_axis_clmat :: circline_mat is y_axis_cmat by (simp add: hermitean_def mat_adj_def mat_cnj_def) lift_definition y_axis :: circline is y_axis_clmat done lemma special_points_on_y_axis' [simp]: shows "on_circline y_axis 0\<^sub>h" and "on_circline y_axis ii\<^sub>h" and "on_circline y_axis \\<^sub>h" by (transfer, transfer, simp add: vec_cnj_def)+ lemma special_points_on_y_axis'' [simp]: shows "0\<^sub>h \ circline_set y_axis" and "ii\<^sub>h \ circline_set y_axis" and "\\<^sub>h \ circline_set y_axis" unfolding circline_set_def by auto lemma on_circline_y_axis: shows "on_circline y_axis z \ (\ c. is_imag c \ z = of_complex c) \ z = \\<^sub>h" proof safe fix z c assume "is_imag c" thus "on_circline y_axis (of_complex c)" proof (transfer, transfer) fix c assume "is_imag c" thus "on_circline_cmat_cvec y_axis_cmat (of_complex_cvec c)" using eq_minus_cnj_iff_imag[of c] by (simp add: vec_cnj_def) qed next fix z assume "on_circline y_axis z" "z \ \\<^sub>h" thus "\c. is_imag c \ z = of_complex c" proof (transfer, transfer, safe) fix a b assume "(a, b) \ vec_zero" "on_circline_cmat_cvec y_axis_cmat (a, b)" "\ (a, b) \\<^sub>v \\<^sub>v" hence "b \ 0" "cnj a * b + cnj b * a = 0" using inf_cvec_z2_zero_iff by (blast, smt add.left_neutral add_cancel_right_right mult.commute mult.left_neutral mult_not_zero on_circline_cmat_cvec_circline_equation y_axis_cmat_def) thus "\c. is_imag c \ (a, b) \\<^sub>v of_complex_cvec c" using eq_minus_cnj_iff_imag[of "a / b"] apply (rule_tac x="a/b" in exI) apply (auto simp add: field_simps) apply (rule_tac x="1/b" in exI, simp) using add_eq_0_iff apply blast apply (rule_tac x="1/b" in exI, simp) done qed next show "on_circline y_axis \\<^sub>h" by simp qed lemma on_circline_y_axis_I [simp]: assumes "is_imag z" shows "on_circline y_axis (of_complex z)" using assms unfolding on_circline_y_axis by auto lemma circline_set_y_axis: shows "circline_set y_axis = of_complex ` {x. is_imag x} \ {\\<^sub>h}" using on_circline_y_axis unfolding circline_set_def by auto lemma circline_set_y_axis_I: assumes "is_imag z" shows "of_complex z \ circline_set y_axis" using assms unfolding circline_set_y_axis by auto text \Positive and negative part of y-axis\ definition positive_y_axis where "positive_y_axis = {z. z \ circline_set y_axis \ z \ \\<^sub>h \ Im (to_complex z) > 0}" definition negative_y_axis where "negative_y_axis = {z. z \ circline_set y_axis \ z \ \\<^sub>h \ Im (to_complex z) < 0}" lemma circline_set_positive_y_axis_I [simp]: assumes "is_imag z" and "Im z > 0" shows "of_complex z \ positive_y_axis" using assms unfolding positive_y_axis_def by simp lemma circline_set_negative_y_axis_I [simp]: assumes "is_imag z" and "Im z < 0" shows "of_complex z \ negative_y_axis" using assms unfolding negative_y_axis_def by simp (* ---------------------------------------------------------------------------- *) subsubsection \Point zero as a circline\ (* ---------------------------------------------------------------------------- *) definition circline_point_0_cmat :: complex_mat where [simp]: "circline_point_0_cmat = (1, 0, 0, 0)" lift_definition circline_point_0_clmat :: circline_mat is circline_point_0_cmat by (simp add: hermitean_def mat_adj_def mat_cnj_def) lift_definition circline_point_0 :: circline is circline_point_0_clmat done lemma circline_type_circline_point_0 [simp]: shows "circline_type circline_point_0 = 0" by (transfer, transfer, simp) lemma zero_in_circline_point_0 [simp]: shows "0\<^sub>h \ circline_set circline_point_0" unfolding circline_set_def by auto (transfer, transfer, simp add: vec_cnj_def)+ (* ---------------------------------------------------------------------------- *) subsubsection \Imaginary unit circle\ (* ---------------------------------------------------------------------------- *) definition imag_unit_circle_cmat :: complex_mat where [simp]: "imag_unit_circle_cmat = (1, 0, 0, 1)" lift_definition imag_unit_circle_clmat :: circline_mat is imag_unit_circle_cmat by (simp add: hermitean_def mat_adj_def mat_cnj_def) lift_definition imag_unit_circle :: circline is imag_unit_circle_clmat done lemma circline_type_imag_unit_circle [simp]: shows "circline_type imag_unit_circle = 1" by (transfer, transfer, simp) (* ----------------------------------------------------------------- *) subsection \Intersection of circlines\ (* ----------------------------------------------------------------- *) definition circline_intersection :: "circline \ circline \ complex_homo set" where "circline_intersection H1 H2 = {z. on_circline H1 z \ on_circline H2 z}" lemma circline_equation_cancel_z2: assumes "circline_equation A B C D z1 z2 " and "z2 \ 0" shows "circline_equation A B C D (z1/z2) 1" using assms by (simp add: field_simps) lemma circline_equation_quadratic_equation: assumes "circline_equation A B (cnj B) D z 1" and "Re z = x" and "Im z = y" and "Re B = bx" and "Im B = by" shows "A*x\<^sup>2 + A*y\<^sup>2 + 2*bx*x + 2*by*y + D = 0" using assms proof- have "z = x + \*y" "B = bx + \*by" using assms complex_eq by auto thus ?thesis using assms by (simp add: field_simps power2_eq_square) qed lemma circline_intersection_symetry: shows "circline_intersection H1 H2 = circline_intersection H2 H1" unfolding circline_intersection_def by auto (* ----------------------------------------------------------------- *) subsection \Möbius action on circlines\ (* ----------------------------------------------------------------- *) definition moebius_circline_cmat_cmat :: "complex_mat \ complex_mat \ complex_mat" where [simp]: "moebius_circline_cmat_cmat M H = congruence (mat_inv M) H" lift_definition moebius_circline_mmat_clmat :: "moebius_mat \ circline_mat \ circline_mat" is moebius_circline_cmat_cmat using mat_det_inv congruence_nonzero hermitean_congruence by simp lift_definition moebius_circline :: "moebius \ circline \ circline" is moebius_circline_mmat_clmat proof transfer fix M M' H H' assume "moebius_cmat_eq M M'" "circline_eq_cmat H H'" thus "circline_eq_cmat (moebius_circline_cmat_cmat M H) (moebius_circline_cmat_cmat M' H')" by (auto simp add: mat_inv_mult_sm) (rule_tac x="ka / Re (k * cnj k)" in exI, auto simp add: complex_mult_cnj_cmod power2_eq_square) qed lemma moebius_preserve_circline_type [simp]: shows "circline_type (moebius_circline M H) = circline_type H" proof (transfer, transfer) fix M H :: complex_mat assume "mat_det M \ 0" "hermitean H \ H \ mat_zero" thus "circline_type_cmat (moebius_circline_cmat_cmat M H) = circline_type_cmat H" using Re_det_sgn_congruence[of "mat_inv M" "H"] mat_det_inv[of "M"] by (simp del: congruence_def) qed text \The central lemma in this section connects the action of Möbius transformations on points and on circlines.\ lemma moebius_circline: shows "{z. on_circline (moebius_circline M H) z} = moebius_pt M ` {z. on_circline H z}" proof safe fix z assume "on_circline H z" thus "on_circline (moebius_circline M H) (moebius_pt M z)" proof (transfer, transfer) fix z :: complex_vec and M H :: complex_mat assume hh: "hermitean H \ H \ mat_zero" "z \ vec_zero" "mat_det M \ 0" let ?z = "M *\<^sub>m\<^sub>v z" let ?H = "mat_adj (mat_inv M) *\<^sub>m\<^sub>m H *\<^sub>m\<^sub>m (mat_inv M)" assume *: "on_circline_cmat_cvec H z" hence "quad_form z H = 0" by simp hence "quad_form ?z ?H = 0" using quad_form_congruence[of M z H] hh by simp thus "on_circline_cmat_cvec (moebius_circline_cmat_cmat M H) (moebius_pt_cmat_cvec M z)" by simp qed next fix z assume "on_circline (moebius_circline M H) z" hence "\ z'. z = moebius_pt M z' \ on_circline H z'" proof (transfer, transfer) fix z :: complex_vec and M H :: complex_mat assume hh: "hermitean H \ H \ mat_zero" "z \ vec_zero" "mat_det M \ 0" let ?iM = "mat_inv M" let ?z' = "?iM *\<^sub>m\<^sub>v z" assume *: "on_circline_cmat_cvec (moebius_circline_cmat_cmat M H) z" have "?z' \ vec_zero" using hh using mat_det_inv mult_mv_nonzero by auto moreover have "z \\<^sub>v moebius_pt_cmat_cvec M ?z'" using hh eye_mv_l mat_inv_r by simp moreover have "M *\<^sub>m\<^sub>v (?iM *\<^sub>m\<^sub>v z) = z" using hh eye_mv_l mat_inv_r by auto hence "on_circline_cmat_cvec H ?z'" using hh * using quad_form_congruence[of M "?iM *\<^sub>m\<^sub>v z" H, symmetric] unfolding moebius_circline_cmat_cmat_def unfolding on_circline_cmat_cvec_def by simp ultimately show "\z'\{v. v \ vec_zero}. z \\<^sub>v moebius_pt_cmat_cvec M z' \ on_circline_cmat_cvec H z'" by blast qed thus "z \ moebius_pt M ` {z. on_circline H z}" by auto qed lemma on_circline_moebius_circline_I [simp]: assumes "on_circline H z" shows "on_circline (moebius_circline M H) (moebius_pt M z)" using assms moebius_circline by fastforce lemma circline_set_moebius_circline [simp]: shows "circline_set (moebius_circline M H) = moebius_pt M ` circline_set H" using moebius_circline[of M H] unfolding circline_set_def by auto lemma circline_set_moebius_circline_I [simp]: assumes "z \ circline_set H" shows "moebius_pt M z \ circline_set (moebius_circline M H)" using assms by simp lemma circline_set_moebius_circline_E: assumes "moebius_pt M z \ circline_set (moebius_circline M H)" shows "z \ circline_set H" using assms using moebius_pt_eq_I[of M z] by auto lemma circline_set_moebius_circline_iff [simp]: shows "moebius_pt M z \ circline_set (moebius_circline M H) \ z \ circline_set H" using moebius_pt_eq_I[of M z] by auto lemma inj_moebius_circline: shows "inj (moebius_circline M)" unfolding inj_on_def proof (safe) fix H H' assume "moebius_circline M H = moebius_circline M H'" thus "H = H'" proof (transfer, transfer) fix M H H' :: complex_mat assume hh: "mat_det M \ 0" let ?iM = "mat_inv M" assume "circline_eq_cmat (moebius_circline_cmat_cmat M H) (moebius_circline_cmat_cmat M H')" then obtain k where "congruence ?iM H' = congruence ?iM (cor k *\<^sub>s\<^sub>m H)" "k \ 0" by auto thus "circline_eq_cmat H H'" using hh inj_congruence[of ?iM H' "cor k *\<^sub>s\<^sub>m H"] mat_det_inv[of M] by auto qed qed lemma moebius_circline_eq_I: assumes "moebius_circline M H1 = moebius_circline M H2" shows "H1 = H2" using assms inj_moebius_circline[of M] unfolding inj_on_def by blast lemma moebius_circline_neq_I [simp]: assumes "H1 \ H2" shows "moebius_circline M H1 \ moebius_circline M H2" using assms inj_moebius_circline[of M] unfolding inj_on_def by blast (* ---------------------------------------------------------------------------- *) subsubsection \Group properties of Möbius action on ciclines\ (* ---------------------------------------------------------------------------- *) text \Möbius actions on circlines have similar properties as Möbius actions on points.\ lemma moebius_circline_id [simp]: shows "moebius_circline id_moebius H = H" by (transfer, transfer) (simp add: mat_adj_def mat_cnj_def, rule_tac x=1 in exI, auto) lemma moebius_circline_comp [simp]: shows "moebius_circline (moebius_comp M1 M2) H = moebius_circline M1 (moebius_circline M2 H)" by (transfer, transfer) (simp add: mat_inv_mult_mm, rule_tac x=1 in exI, simp add: mult_mm_assoc) lemma moebius_circline_comp_inv_left [simp]: shows "moebius_circline (moebius_inv M) (moebius_circline M H) = H" by (subst moebius_circline_comp[symmetric], simp) lemma moebius_circline_comp_inv_right [simp]: shows "moebius_circline M (moebius_circline (moebius_inv M) H) = H" by (subst moebius_circline_comp[symmetric], simp) (* ----------------------------------------------------------------- *) subsection \Action of Euclidean similarities on circlines\ (* ----------------------------------------------------------------- *) lemma moebius_similarity_lines_to_lines [simp]: assumes "a \ 0" shows "\\<^sub>h \ circline_set (moebius_circline (moebius_similarity a b) H) \ \\<^sub>h \ circline_set H" using assms by (metis circline_set_moebius_circline_iff moebius_similarity_inf) lemma moebius_similarity_lines_to_lines': assumes "a \ 0" shows "on_circline (moebius_circline (moebius_similarity a b) H) \\<^sub>h \ \\<^sub>h \ circline_set H" using moebius_similarity_lines_to_lines assms unfolding circline_set_def by simp (* ----------------------------------------------------------------- *) subsection \Conjugation, recpiprocation and inversion of circlines\ (* ----------------------------------------------------------------- *) text \Conjugation of circlines\ definition conjugate_circline_cmat :: "complex_mat \ complex_mat" where [simp]: "conjugate_circline_cmat = mat_cnj" lift_definition conjugate_circline_clmat :: "circline_mat \ circline_mat" is conjugate_circline_cmat by (auto simp add: hermitean_def mat_adj_def mat_cnj_def) lift_definition conjugate_circline :: "circline \ circline" is conjugate_circline_clmat by transfer (metis circline_eq_cmat_def conjugate_circline_cmat_def hermitean_transpose mat_t_mult_sm) lemma conjugate_circline_set': shows "conjugate ` circline_set H \ circline_set (conjugate_circline H)" proof (safe) fix z assume "z \ circline_set H" thus "conjugate z \ circline_set (conjugate_circline H)" unfolding circline_set_def apply simp apply (transfer, transfer) unfolding on_circline_cmat_cvec_def conjugate_cvec_def conjugate_circline_cmat_def apply (subst quad_form_vec_cnj_mat_cnj, simp_all) done qed lemma conjugate_conjugate_circline [simp]: shows "conjugate_circline (conjugate_circline H) = H" by (transfer, transfer, force) lemma circline_set_conjugate_circline [simp]: shows "circline_set (conjugate_circline H) = conjugate ` circline_set H" (is "?lhs = ?rhs") proof (safe) fix z assume "z \ ?lhs" show "z \ ?rhs" proof show "z = conjugate (conjugate z)" by simp next show "conjugate z \ circline_set H" using \z \ circline_set (conjugate_circline H)\ using conjugate_circline_set'[of "conjugate_circline H"] by auto qed next fix z assume "z \ circline_set H" thus "conjugate z \ circline_set (conjugate_circline H)" using conjugate_circline_set'[of H] by auto qed lemma on_circline_conjugate_circline [simp]: shows "on_circline (conjugate_circline H) z \ on_circline H (conjugate z)" using circline_set_conjugate_circline[of H] unfolding circline_set_def by force text \Inversion of circlines\ definition circline_inversion_cmat :: "complex_mat \ complex_mat" where [simp]: "circline_inversion_cmat H = (let (A, B, C, D) = H in (D, B, C, A))" lift_definition circline_inversion_clmat :: "circline_mat \ circline_mat" is circline_inversion_cmat by (auto simp add: hermitean_def mat_adj_def mat_cnj_def) lift_definition circline_inversion :: "circline \ circline" is circline_inversion_clmat by transfer auto lemma on_circline_circline_inversion [simp]: shows "on_circline (circline_inversion H) z \ on_circline H (reciprocal (conjugate z))" by (transfer, transfer, auto simp add: vec_cnj_def field_simps) lemma circline_set_circline_inversion [simp]: shows "circline_set (circline_inversion H) = inversion ` circline_set H" unfolding circline_set_def inversion_def by (force simp add: comp_def image_iff) text \Reciprocal of circlines\ definition circline_reciprocal :: "circline \ circline" where "circline_reciprocal = conjugate_circline \ circline_inversion" lemma circline_set_circline_reciprocal: shows "circline_set (circline_reciprocal H) = reciprocal ` circline_set H" unfolding circline_reciprocal_def comp_def by (auto simp add: inversion_def image_iff) text \Rotation of circlines\ lemma rotation_pi_2_y_axis [simp]: shows "moebius_circline (moebius_rotation (pi/2)) y_axis = x_axis" unfolding moebius_rotation_def moebius_similarity_def by (transfer, transfer, simp add: mat_adj_def mat_cnj_def) lemma rotation_minus_pi_2_y_axis [simp]: shows "moebius_circline (moebius_rotation (-pi/2)) y_axis = x_axis" unfolding moebius_rotation_def moebius_similarity_def by (transfer, transfer, simp add: mat_adj_def mat_cnj_def, rule_tac x="-1" in exI, simp) lemma rotation_minus_pi_2_x_axis [simp]: shows "moebius_circline (moebius_rotation (-pi/2)) x_axis = y_axis" unfolding moebius_rotation_def moebius_similarity_def by (transfer, transfer, simp add: mat_adj_def mat_cnj_def) lemma rotation_pi_2_x_axis [simp]: shows "moebius_circline (moebius_rotation (pi/2)) x_axis = y_axis" unfolding moebius_rotation_def moebius_similarity_def by (transfer, transfer, simp add: mat_adj_def mat_cnj_def, rule_tac x="-1" in exI, simp) lemma rotation_minus_pi_2_positive_y_axis [simp]: shows "(moebius_pt (moebius_rotation (-pi/2))) ` positive_y_axis = positive_x_axis" proof safe fix y assume y: "y \ positive_y_axis" have *: "Re (a * \ / b) < 0 \ Im (a / b) > 0" for a b by (subst times_divide_eq_left [symmetric], subst mult.commute, subst Re_i_times) auto from y * show "moebius_pt (moebius_rotation (-pi/2)) y \ positive_x_axis" unfolding positive_y_axis_def positive_x_axis_def circline_set_def unfolding moebius_rotation_def moebius_similarity_def apply simp apply transfer apply transfer apply (auto simp add: vec_cnj_def field_simps add_eq_0_iff) done next fix x assume x: "x \ positive_x_axis" let ?y = "moebius_pt (moebius_rotation (pi/2)) x" have *: "Im (a * \ / b) > 0 \ Re (a / b) > 0" for a b by (subst times_divide_eq_left [symmetric], subst mult.commute, subst Im_i_times) auto hence "?y \ positive_y_axis" using \x \ positive_x_axis\ unfolding positive_x_axis_def positive_y_axis_def unfolding moebius_rotation_def moebius_similarity_def unfolding circline_set_def apply simp apply transfer apply transfer apply (auto simp add: vec_cnj_def field_simps add_eq_0_iff) done thus "x \ moebius_pt (moebius_rotation (-pi/2)) ` positive_y_axis" by (auto simp add: image_iff) (rule_tac x="?y" in bexI, simp_all) qed (* ----------------------------------------------------------------- *) subsection \Circline uniqueness\ (* ----------------------------------------------------------------- *) (* ----------------------------------------------------------------- *) subsubsection \Zero type circline uniqueness\ (* ----------------------------------------------------------------- *) lemma unique_circline_type_zero_0': shows "(circline_type circline_point_0 = 0 \ 0\<^sub>h \ circline_set circline_point_0) \ (\ H. circline_type H = 0 \ 0\<^sub>h \ circline_set H \ H = circline_point_0)" unfolding circline_set_def proof (safe) show "circline_type circline_point_0 = 0" by (transfer, transfer, simp) next show "on_circline circline_point_0 0\<^sub>h" using circline_set_def zero_in_circline_point_0 by auto next fix H assume "circline_type H = 0" "on_circline H 0\<^sub>h" thus "H = circline_point_0" proof (transfer, transfer) fix H :: complex_mat assume hh: "hermitean H \ H \ mat_zero" obtain A B C D where HH: "H = (A, B, C, D)" by (cases "H") auto hence *: "C = cnj B" "is_real A" using hh hermitean_elems[of A B C D] by auto assume "circline_type_cmat H = 0" "on_circline_cmat_cvec H 0\<^sub>v" thus "circline_eq_cmat H circline_point_0_cmat" using HH hh * by (simp add: Let_def vec_cnj_def sgn_minus sgn_mult sgn_zero_iff) (rule_tac x="1/Re A" in exI, cases A, cases B, simp add: Complex_eq sgn_zero_iff) qed qed lemma unique_circline_type_zero_0: shows "\! H. circline_type H = 0 \ 0\<^sub>h \ circline_set H" using unique_circline_type_zero_0' by blast lemma unique_circline_type_zero: shows "\! H. circline_type H = 0 \ z \ circline_set H" proof- obtain M where ++: "moebius_pt M z = 0\<^sub>h" using ex_moebius_1[of z] by auto have +++: "z = moebius_pt (moebius_inv M) 0\<^sub>h" by (subst ++[symmetric]) simp then obtain H0 where *: "circline_type H0 = 0 \ 0\<^sub>h \ circline_set H0" and **: "\ H'. circline_type H' = 0 \ 0\<^sub>h \ circline_set H' \ H' = H0" using unique_circline_type_zero_0 by auto let ?H' = "moebius_circline (moebius_inv M) H0" show ?thesis unfolding Ex1_def using * +++ proof (rule_tac x="?H'" in exI, simp, safe) fix H' assume "circline_type H' = 0" "moebius_pt (moebius_inv M) 0\<^sub>h \ circline_set H'" hence "0\<^sub>h \ circline_set (moebius_circline M H')" using ++ +++ by force hence "moebius_circline M H' = H0" using **[rule_format, of "moebius_circline M H'"] using \circline_type H' = 0\ by simp thus "H' = moebius_circline (moebius_inv M) H0" by auto qed qed (* ----------------------------------------------------------------- *) subsubsection \Negative type circline uniqueness\ (* ----------------------------------------------------------------- *) lemma unique_circline_01inf': shows "0\<^sub>h \ circline_set x_axis \ 1\<^sub>h \ circline_set x_axis \ \\<^sub>h \ circline_set x_axis \ (\ H. 0\<^sub>h \ circline_set H \ 1\<^sub>h \ circline_set H \ \\<^sub>h \ circline_set H \ H = x_axis)" proof safe fix H assume "0\<^sub>h \ circline_set H" "1\<^sub>h \ circline_set H" "\\<^sub>h \ circline_set H" thus "H = x_axis" unfolding circline_set_def apply simp proof (transfer, transfer) fix H assume hh: "hermitean H \ H \ mat_zero" obtain A B C D where HH: "H = (A, B, C, D)" by (cases H) auto have *: "C = cnj B" "A = 0 \ D = 0 \ B \ 0" using hermitean_elems[of A B C D] hh HH by auto obtain Bx By where "B = Complex Bx By" by (cases B) auto assume "on_circline_cmat_cvec H 0\<^sub>v" "on_circline_cmat_cvec H 1\<^sub>v" "on_circline_cmat_cvec H \\<^sub>v" thus "circline_eq_cmat H x_axis_cmat" using * HH \C = cnj B\ \B = Complex Bx By\ by (simp add: Let_def vec_cnj_def Complex_eq) (rule_tac x="1/By" in exI, auto) qed qed simp_all lemma unique_circline_set: assumes "A \ B" and "A \ C" and "B \ C" shows "\! H. A \ circline_set H \ B \ circline_set H \ C \ circline_set H" proof- let ?P = "\ A B C. A \ B \ A \ C \ B \ C \ (\! H. A \ circline_set H \ B \ circline_set H \ C \ circline_set H)" have "?P A B C" proof (rule wlog_moebius_01inf[of ?P]) fix M a b c let ?M = "moebius_pt M" assume "?P a b c" show "?P (?M a) (?M b) (?M c)" proof assume "?M a \ ?M b \ ?M a \ ?M c \ ?M b \ ?M c" hence "a \ b" "b \ c" "a \ c" by auto hence "\!H. a \ circline_set H \ b \ circline_set H \ c \ circline_set H" using \?P a b c\ by simp then obtain H where *: "a \ circline_set H \ b \ circline_set H \ c \ circline_set H" and **: "\H'. a \ circline_set H' \ b \ circline_set H' \ c \ circline_set H' \ H' = H" unfolding Ex1_def by auto let ?H' = "moebius_circline M H" show "\! H. ?M a \ circline_set H \ moebius_pt M b \ circline_set H \ moebius_pt M c \ circline_set H" unfolding Ex1_def proof (rule_tac x="?H'" in exI, rule) show "?M a \ circline_set ?H' \ ?M b \ circline_set ?H' \ ?M c \ circline_set ?H'" using * by auto next show "\H'. ?M a \ circline_set H' \ ?M b \ circline_set H' \ ?M c \ circline_set H' \ H' = ?H'" proof (safe) fix H' let ?iH' = "moebius_circline (moebius_inv M) H'" assume "?M a \ circline_set H'" "?M b \ circline_set H'" "?M c \ circline_set H'" hence "a \ circline_set ?iH' \ b \ circline_set ?iH' \ c \ circline_set ?iH'" by simp hence "H = ?iH'" using ** by blast thus "H' = moebius_circline M H" by simp qed qed qed next show "?P 0\<^sub>h 1\<^sub>h \\<^sub>h" using unique_circline_01inf' unfolding Ex1_def by (safe, rule_tac x="x_axis" in exI) auto qed fact+ thus ?thesis using assms by simp qed lemma zero_one_inf_x_axis [simp]: assumes "0\<^sub>h \ circline_set H" and "1\<^sub>h \ circline_set H" and "\\<^sub>h \ circline_set H" shows "H = x_axis" using assms unique_circline_set[of "0\<^sub>h" "1\<^sub>h" "\\<^sub>h"] by auto (* ----------------------------------------------------------------- *) subsection \Circline set cardinality\ (* ----------------------------------------------------------------- *) (* ----------------------------------------------------------------- *) subsubsection \Diagonal circlines\ (* ----------------------------------------------------------------- *) definition is_diag_circline_cmat :: "complex_mat \ bool" where [simp]: "is_diag_circline_cmat H = (let (A, B, C, D) = H in B = 0 \ C = 0)" lift_definition is_diag_circline_clmat :: "circline_mat \ bool" is is_diag_circline_cmat done lift_definition circline_diag :: "circline \ bool" is is_diag_circline_clmat by transfer auto lemma circline_diagonalize: shows "\ M H'. moebius_circline M H = H' \ circline_diag H'" proof (transfer, transfer) fix H assume hh: "hermitean H \ H \ mat_zero" obtain A B C D where HH: "H = (A, B, C, D)" by (cases "H") auto hence HH_elems: "is_real A" "is_real D" "C = cnj B" using hermitean_elems[of A B C D] hh by auto obtain M k1 k2 where *: "mat_det M \ 0" "unitary M" "congruence M H = (k1, 0, 0, k2)" "is_real k1" "is_real k2" using hermitean_diagonizable[of H] hh by auto have "k1 \ 0 \ k2 \ 0" using \congruence M H = (k1, 0, 0, k2)\ hh congruence_nonzero[of H M] \mat_det M \ 0\ by auto let ?M' = "mat_inv M" let ?H' = "(k1, 0, 0, k2)" have "circline_eq_cmat (moebius_circline_cmat_cmat ?M' H) ?H' \ is_diag_circline_cmat ?H'" using * by force moreover have "?H' \ hermitean_nonzero" using * \k1 \ 0 \ k2 \ 0\ eq_cnj_iff_real[of k1] eq_cnj_iff_real[of k2] by (auto simp add: hermitean_def mat_adj_def mat_cnj_def) moreover have "mat_det ?M' \ 0" using * mat_det_inv[of M] by auto ultimately show "\M\{M. mat_det M \ 0}. \H'\hermitean_nonzero. circline_eq_cmat (moebius_circline_cmat_cmat M H) H' \ is_diag_circline_cmat H'" by blast qed lemma wlog_circline_diag: assumes "\ H. circline_diag H \ P H" "\ M H. P H \ P (moebius_circline M H)" shows "P H" proof- obtain M H' where "moebius_circline M H = H'" "circline_diag H'" using circline_diagonalize[of H] by auto hence "P (moebius_circline M H)" using assms(1) by simp thus ?thesis using assms(2)[of "moebius_circline M H" "moebius_inv M"] by simp qed (* ----------------------------------------------------------------- *) subsubsection \Zero type circline set cardinality\ (* ----------------------------------------------------------------- *) lemma circline_type_zero_card_eq1_0: assumes "circline_type H = 0" and "0\<^sub>h \ circline_set H" shows "circline_set H = {0\<^sub>h}" using assms unfolding circline_set_def proof(safe) fix z assume "on_circline H z" "circline_type H = 0" "on_circline H 0\<^sub>h" hence "H = circline_point_0" using unique_circline_type_zero_0' unfolding circline_set_def by simp thus "z = 0\<^sub>h" using \on_circline H z\ by (transfer, transfer) (case_tac z, case_tac H, force simp add: vec_cnj_def) qed lemma circline_type_zero_card_eq1: assumes "circline_type H = 0" shows "\ z. circline_set H = {z}" proof- have "\ z. on_circline H z" using assms proof (transfer, transfer) fix H assume hh: "hermitean H \ H \ mat_zero" obtain A B C D where HH: "H = (A, B, C, D)" by (cases H) auto hence "C = cnj B" "is_real A" "is_real D" using hh hermitean_elems[of A B C D] by auto assume "circline_type_cmat H = 0" hence "mat_det H = 0" by (simp add: complex_eq_if_Re_eq hh mat_det_hermitean_real sgn_eq_0_iff) hence "A*D = B*C" using HH by simp show "Bex {v. v \ vec_zero} (on_circline_cmat_cvec H)" proof (cases "A \ 0 \ B \ 0") case True thus ?thesis using HH \A*D = B*C\ by (rule_tac x="(-B, A)" in bexI) (auto simp add: Let_def vec_cnj_def field_simps) next case False thus ?thesis using HH \C = cnj B\ by (rule_tac x="(1, 0)" in bexI) (simp_all add: Let_def vec_cnj_def) qed qed then obtain z where "on_circline H z" by auto obtain M where "moebius_pt M z = 0\<^sub>h" using ex_moebius_1[of z] by auto hence "0\<^sub>h \ circline_set (moebius_circline M H)" using on_circline_moebius_circline_I[OF \on_circline H z\, of M] unfolding circline_set_def by simp hence "circline_set (moebius_circline M H) = {0\<^sub>h}" using circline_type_zero_card_eq1_0[of "moebius_circline M H"] \circline_type H = 0\ by auto hence "circline_set H = {z}" using \moebius_pt M z = 0\<^sub>h\ using bij_moebius_pt[of M] bij_image_singleton[of "moebius_pt M" "circline_set H" _ z] by simp thus ?thesis by auto qed (* ----------------------------------------------------------------- *) subsubsection \Negative type circline set cardinality\ (* ----------------------------------------------------------------- *) lemma quad_form_diagonal_iff: assumes "k1 \ 0" and "is_real k1" and "is_real k2" and "Re k1 * Re k2 < 0" shows "quad_form (z1, 1) (k1, 0, 0, k2) = 0 \ (\ \. z1 = rcis (sqrt (Re (-k2 /k1))) \)" proof- have "Re (-k2/k1) \ 0" using \Re k1 * Re k2 < 0\ \is_real k1\ \is_real k2\ \k1 \ 0\ using Re_divide_real[of k1 "-k2"] by (smt divide_less_0_iff mult_nonneg_nonneg mult_nonpos_nonpos uminus_complex.simps(1)) have "quad_form (z1, 1) (k1, 0, 0, k2) = 0 \ (cor (cmod z1))\<^sup>2 = -k2 / k1" using assms add_eq_0_iff[of k2 "k1*(cor (cmod z1))\<^sup>2"] using eq_divide_imp[of k1 "(cor (cmod z1))\<^sup>2" "-k2"] by (auto simp add: vec_cnj_def field_simps complex_mult_cnj_cmod) also have "... \ (cmod z1)\<^sup>2 = Re (-k2 /k1)" using assms apply (subst complex_eq_if_Re_eq) using Re_complex_of_real[of "(cmod z1)\<^sup>2"] div_reals by auto also have "... \ cmod z1 = sqrt (Re (-k2 /k1))" by (metis norm_ge_zero real_sqrt_ge_0_iff real_sqrt_pow2 real_sqrt_power) also have "... \ (\ \. z1 = rcis (sqrt (Re (-k2 /k1))) \)" using rcis_cmod_arg[of z1, symmetric] assms abs_of_nonneg[of "sqrt (Re (-k2/k1))"] using \Re (-k2/k1) \ 0\ by auto finally show ?thesis . qed lemma circline_type_neg_card_gt3_diag: assumes "circline_type H < 0" and "circline_diag H" shows "\ A B C. A \ B \ A \ C \ B \ C \ {A, B, C} \ circline_set H" using assms unfolding circline_set_def apply (simp del: HOL.ex_simps) proof (transfer, transfer) fix H assume hh: "hermitean H \ H \ mat_zero" obtain A B C D where HH: "H = (A, B, C, D)" by (cases H) auto hence HH_elems: "is_real A" "is_real D" "C = cnj B" using hermitean_elems[of A B C D] hh by auto assume "circline_type_cmat H < 0" "is_diag_circline_cmat H" hence "B = 0" "C = 0" "Re A * Re D < 0" "A \ 0" using HH \is_real A\ \is_real D\ by auto let ?x = "sqrt (Re (- D / A))" let ?A = "(rcis ?x 0, 1)" let ?B = "(rcis ?x (pi/2), 1)" let ?C = "(rcis ?x pi, 1)" from quad_form_diagonal_iff[OF \A \ 0\ \is_real A\ \is_real D\ \Re A * Re D < 0\] have "quad_form ?A (A, 0, 0, D) = 0" "quad_form ?B (A, 0, 0, D) = 0" "quad_form ?C (A, 0, 0, D) = 0" by (auto simp del: rcis_zero_arg) hence "on_circline_cmat_cvec H ?A \ on_circline_cmat_cvec H ?B \ on_circline_cmat_cvec H ?C" using HH \B = 0\ \C = 0\ by simp moreover have "Re (D / A) < 0" using \Re A * Re D < 0\ \A \ 0\ \is_real A\ \is_real D\ using Re_divide_real[of A D] by (metis Re_complex_div_lt_0 Re_mult_real div_reals eq_cnj_iff_real is_real_div) hence "\ ?A \\<^sub>v ?B \ \ ?A \\<^sub>v ?C \ \ ?B \\<^sub>v ?C" unfolding rcis_def by (auto simp add: cis_def complex.corec) moreover have "?A \ vec_zero" "?B \ vec_zero" "?C \ vec_zero" by auto ultimately show "\A\{v. v \ vec_zero}. \B\{v. v \ vec_zero}. \C\{v. v \ vec_zero}. \ A \\<^sub>v B \ \ A \\<^sub>v C \ \ B \\<^sub>v C \ on_circline_cmat_cvec H A \ on_circline_cmat_cvec H B \ on_circline_cmat_cvec H C" by blast qed lemma circline_type_neg_card_gt3: assumes "circline_type H < 0" shows "\ A B C. A \ B \ A \ C \ B \ C \ {A, B, C} \ circline_set H" proof- obtain M H' where "moebius_circline M H = H'" "circline_diag H'" using circline_diagonalize[of H] assms by auto moreover hence "circline_type H' < 0" using assms moebius_preserve_circline_type by auto ultimately obtain A B C where "A \ B" "A \ C" "B \ C" "{A, B, C} \ circline_set H'" using circline_type_neg_card_gt3_diag[of H'] by auto let ?iM = "moebius_inv M" have "moebius_circline ?iM H' = H" using \moebius_circline M H = H'\[symmetric] by simp let ?A = "moebius_pt ?iM A" and ?B= "moebius_pt ?iM B" and ?C = "moebius_pt ?iM C" have "?A \ circline_set H" "?B \ circline_set H" "?C \ circline_set H" using \moebius_circline ?iM H' = H\[symmetric] \{A, B, C} \ circline_set H'\ by simp_all moreover have "?A \ ?B" "?A \ ?C" "?B \ ?C" using \A \ B\ \A \ C\ \B \ C\ by auto ultimately show ?thesis by auto qed (* ----------------------------------------------------------------- *) subsubsection \Positive type circline set cardinality\ (* ----------------------------------------------------------------- *) lemma circline_type_pos_card_eq0_diag: assumes "circline_diag H" and "circline_type H > 0" shows "circline_set H = {}" using assms unfolding circline_set_def apply simp proof (transfer, transfer) fix H assume hh: "hermitean H \ H \ mat_zero" obtain A B C D where HH: "H = (A, B, C, D)" by (cases H) auto hence HH_elems: "is_real A" "is_real D" "C = cnj B" using hermitean_elems[of A B C D] hh by auto assume "is_diag_circline_cmat H" "0 < circline_type_cmat H" hence "B = 0" "C = 0" "Re A * Re D > 0" "A \ 0" using HH \is_real A\ \is_real D\ by auto show "\x\{v. v \ vec_zero}. \ on_circline_cmat_cvec H x" proof fix x assume "x \ {v. v \ vec_zero}" obtain x1 x2 where xx: "x = (x1, x2)" by (cases x, auto) have "(Re A > 0 \ Re D > 0) \ (Re A < 0 \ Re D < 0)" using \Re A * Re D > 0\ by (metis linorder_neqE_linordered_idom mult_eq_0_iff zero_less_mult_pos zero_less_mult_pos2) moreover have "(Re (x1 * cnj x1) \ 0 \ Re (x2 * cnj x2) > 0) \ (Re (x1 * cnj x1) > 0 \ Re (x2 * cnj x2) \ 0)" using \x \ {v. v \ vec_zero}\ xx apply auto apply (simp add: complex_neq_0 power2_eq_square)+ done ultimately have "Re A * Re (x1 * cnj x1) + Re D * Re (x2 * cnj x2) \ 0" by (smt mult_neg_pos mult_nonneg_nonneg mult_nonpos_nonneg mult_pos_pos) hence "A * (x1 * cnj x1) + D * (x2 * cnj x2) \ 0" using \is_real A\ \is_real D\ by (metis Re_mult_real plus_complex.simps(1) zero_complex.simps(1)) thus "\ on_circline_cmat_cvec H x" using HH \B = 0\ \C = 0\ xx by (simp add: vec_cnj_def field_simps) qed qed lemma circline_type_pos_card_eq0: assumes "circline_type H > 0" shows "circline_set H = {}" proof- obtain M H' where "moebius_circline M H = H'" "circline_diag H'" using circline_diagonalize[of H] assms by auto moreover hence "circline_type H' > 0" using assms moebius_preserve_circline_type by auto ultimately have "circline_set H' = {}" using circline_type_pos_card_eq0_diag[of H'] by auto let ?iM = "moebius_inv M" have "moebius_circline ?iM H' = H" using \moebius_circline M H = H'\[symmetric] by simp thus ?thesis using \circline_set H' = {}\ by auto qed (* ----------------------------------------------------------------- *) subsubsection \Cardinality determines type\ (* ----------------------------------------------------------------- *) lemma card_eq1_circline_type_zero: assumes "\ z. circline_set H = {z}" shows "circline_type H = 0" proof (cases "circline_type H < 0") case True thus ?thesis using circline_type_neg_card_gt3[of H] assms by auto next case False show ?thesis proof (cases "circline_type H > 0") case True thus ?thesis using circline_type_pos_card_eq0[of H] assms by auto next case False thus ?thesis using \\ (circline_type H) < 0\ by simp qed qed (* ----------------------------------------------------------------- *) subsubsection \Circline set is injective\ (* ----------------------------------------------------------------- *) lemma inj_circline_set: assumes "circline_set H = circline_set H'" and "circline_set H \ {}" shows "H = H'" proof (cases "circline_type H < 0") case True then obtain A B C where "A \ B" "A \ C" "B \ C" "{A, B, C} \ circline_set H" using circline_type_neg_card_gt3[of H] by auto hence "\!H. A \ circline_set H \ B \ circline_set H \ C \ circline_set H" using unique_circline_set[of A B C] by simp thus ?thesis using \circline_set H = circline_set H'\ \{A, B, C} \ circline_set H\ by auto next case False show ?thesis proof (cases "circline_type H = 0") case True moreover then obtain A where "{A} = circline_set H" using circline_type_zero_card_eq1[of H] by auto moreover hence "circline_type H' = 0" using \circline_set H = circline_set H'\ card_eq1_circline_type_zero[of H'] by auto ultimately show ?thesis using unique_circline_type_zero[of A] \circline_set H = circline_set H'\ by auto next case False hence "circline_type H > 0" using \\ (circline_type H < 0)\ by auto thus ?thesis using \circline_set H \ {}\ circline_type_pos_card_eq0[of H] by auto qed qed (* ----------------------------------------------------------------- *) subsection \Circline points - cross ratio real\ (* ----------------------------------------------------------------- *) lemma four_points_on_circline_iff_cross_ratio_real: assumes "distinct [z, u, v, w]" shows "is_real (to_complex (cross_ratio z u v w)) \ (\ H. {z, u, v, w} \ circline_set H)" proof- have "\ z. distinct [z, u, v, w] \ is_real (to_complex (cross_ratio z u v w)) \ (\ H. {z, u, v, w} \ circline_set H)" (is "?P u v w") proof (rule wlog_moebius_01inf[of ?P u v w]) fix M a b c assume aa: "?P a b c" let ?Ma = "moebius_pt M a" and ?Mb = "moebius_pt M b" and ?Mc = "moebius_pt M c" show "?P ?Ma ?Mb ?Mc" proof (rule allI, rule impI) fix z obtain d where *: "z = moebius_pt M d" using bij_moebius_pt[of M] unfolding bij_def by auto let ?Md = "moebius_pt M d" assume "distinct [z, moebius_pt M a, moebius_pt M b, moebius_pt M c]" hence "distinct [a, b, c, d]" using * by auto moreover have "(\ H. {d, a, b, c} \ circline_set H) \ (\ H. {z, ?Ma, ?Mb, ?Mc} \ circline_set H)" using * apply auto apply (rule_tac x="moebius_circline M H" in exI, simp) apply (rule_tac x="moebius_circline (moebius_inv M) H" in exI, simp) done ultimately show "is_real (to_complex (cross_ratio z ?Ma ?Mb ?Mc)) = (\H. {z, ?Ma, ?Mb, ?Mc} \ circline_set H)" using aa[rule_format, of d] * by auto qed next show "?P 0\<^sub>h 1\<^sub>h \\<^sub>h" proof safe fix z assume "distinct [z, 0\<^sub>h, 1\<^sub>h, \\<^sub>h]" hence "z \ \\<^sub>h" by auto assume "is_real (to_complex (cross_ratio z 0\<^sub>h 1\<^sub>h \\<^sub>h))" hence "is_real (to_complex z)" by simp hence "z \ circline_set x_axis" using of_complex_to_complex[symmetric, OF \z \ \\<^sub>h\] using circline_set_x_axis by auto thus "\H. {z, 0\<^sub>h, 1\<^sub>h, \\<^sub>h} \ circline_set H" by (rule_tac x=x_axis in exI, auto) next fix z H assume *: "distinct [z, 0\<^sub>h, 1\<^sub>h, \\<^sub>h]" "{z, 0\<^sub>h, 1\<^sub>h, \\<^sub>h} \ circline_set H" hence "H = x_axis" by auto hence "z \ circline_set x_axis" using * by auto hence "is_real (to_complex z)" using * circline_set_x_axis by auto thus "is_real (to_complex (cross_ratio z 0\<^sub>h 1\<^sub>h \\<^sub>h))" by simp qed next show "u \ v" "v \ w" "u \ w" using assms by auto qed thus ?thesis using assms by auto qed (* ----------------------------------------------------------------- *) subsection \Symmetric points wrt. circline\ (* ----------------------------------------------------------------- *) text \In the extended complex plane there are no substantial differences between circles and lines, so we will consider only one kind of relation and call two points \emph{circline symmetric} if they are mapped to one another using either reflection or inversion over arbitrary line or circle. Points are symmetric iff the bilinear form of their representation vectors and matrix is zero.\ definition circline_symmetric_cvec_cmat :: "complex_vec \ complex_vec \ complex_mat \ bool" where [simp]: "circline_symmetric_cvec_cmat z1 z2 H \ bilinear_form z1 z2 H = 0" lift_definition circline_symmetric_hcoords_clmat :: "complex_homo_coords \ complex_homo_coords \ circline_mat \ bool" is circline_symmetric_cvec_cmat done lift_definition circline_symmetric :: "complex_homo \ complex_homo \ circline \ bool" is circline_symmetric_hcoords_clmat apply transfer apply (simp del: bilinear_form_def) apply (erule exE)+ apply (simp add: bilinear_form_scale_m bilinear_form_scale_v1 bilinear_form_scale_v2 del: vec_cnj_sv quad_form_def bilinear_form_def) done lemma symmetry_principle [simp]: assumes "circline_symmetric z1 z2 H" shows "circline_symmetric (moebius_pt M z1) (moebius_pt M z2) (moebius_circline M H)" using assms by (transfer, transfer, simp del: bilinear_form_def congruence_def) text \Symmetry wrt. @{term "unit_circle"}\ lemma circline_symmetric_0inf_disc [simp]: shows "circline_symmetric 0\<^sub>h \\<^sub>h unit_circle" by (transfer, transfer, simp add: vec_cnj_def) lemma circline_symmetric_inv_homo_disc [simp]: shows "circline_symmetric a (inversion a) unit_circle" unfolding inversion_def by (transfer, transfer) (case_tac a, auto simp add: vec_cnj_def) lemma circline_symmetric_inv_homo_disc': assumes "circline_symmetric a a' unit_circle" shows "a' = inversion a" unfolding inversion_def using assms proof (transfer, transfer) fix a a' assume vz: "a \ vec_zero" "a' \ vec_zero" obtain a1 a2 where aa: "a = (a1, a2)" by (cases a, auto) obtain a1' a2' where aa': "a' = (a1', a2')" by (cases a', auto) assume *: "circline_symmetric_cvec_cmat a a' unit_circle_cmat" show "a' \\<^sub>v (conjugate_cvec \ reciprocal_cvec) a" proof (cases "a1' = 0") case True thus ?thesis using aa aa' vz * by (auto simp add: vec_cnj_def field_simps) next case False show ?thesis proof (cases "a2 = 0") case True thus ?thesis using \a1' \ 0\ using aa aa' * vz by (simp add: vec_cnj_def field_simps) next case False thus ?thesis using \a1' \ 0\ aa aa' * by (simp add: vec_cnj_def field_simps) (rule_tac x="cnj a2 / a1'" in exI, simp add: field_simps) qed qed qed lemma ex_moebius_circline_x_axis: assumes "circline_type H < 0" shows "\ M. moebius_circline M H = x_axis" proof- obtain A B C where *: "A \ B" "A \ C" "B \ C" "on_circline H A" "on_circline H B" "on_circline H C" using circline_type_neg_card_gt3[OF assms] unfolding circline_set_def by auto then obtain M where "moebius_pt M A = 0\<^sub>h" "moebius_pt M B = 1\<^sub>h" "moebius_pt M C = \\<^sub>h" using ex_moebius_01inf by blast hence "moebius_circline M H = x_axis" using * by (metis circline_set_I circline_set_moebius_circline rev_image_eqI unique_circline_01inf') thus ?thesis by blast qed lemma wlog_circline_x_axis: assumes "circline_type H < 0" assumes "\ M H. P H \ P (moebius_circline M H)" assumes "P x_axis" shows "P H" proof- obtain M where "moebius_circline M H = x_axis" using ex_moebius_circline_x_axis[OF assms(1)] by blast then obtain M' where "moebius_circline M' x_axis = H" by (metis moebius_circline_comp_inv_left) thus ?thesis using assms(2)[of x_axis M'] assms(3) by simp qed lemma circline_intersection_at_most_2_points: assumes "H1 \ H2" shows "finite (circline_intersection H1 H2) \ card (circline_intersection H1 H2) \ 2" proof (rule ccontr) assume "\ ?thesis" hence "infinite (circline_intersection H1 H2) \ card (circline_intersection H1 H2) > 2" by auto hence "\ A B C. A \ B \ B \ C \ A \ C \ {A, B, C} \ circline_intersection H1 H2" proof assume "card (circline_intersection H1 H2) > 2" thus ?thesis using card_geq_3_iff_contains_3_elems[of "circline_intersection H1 H2"] by auto next assume "infinite (circline_intersection H1 H2)" thus ?thesis using infinite_contains_3_elems by blast qed then obtain A B C where "A \ B" "B \ C" "A \ C" "{A, B, C} \ circline_intersection H1 H2" by blast hence "H2 = H1" using circline_intersection_def mem_Collect_eq unique_circline_set by fastforce thus False using assms by simp qed end diff --git a/thys/Complex_Geometry/Hermitean_Matrices.thy b/thys/Complex_Geometry/Hermitean_Matrices.thy --- a/thys/Complex_Geometry/Hermitean_Matrices.thy +++ b/thys/Complex_Geometry/Hermitean_Matrices.thy @@ -1,418 +1,418 @@ (* -------------------------------------------------------------------------- *) subsection \Hermitean matrices\ (* -------------------------------------------------------------------------- *) text \Hermitean matrices over $\mathbb{C}$ generalize symmetric matrices over $\mathbb{R}$. Quadratic forms with Hermitean matrices represent circles and lines in the extended complex plane (when applied to homogenous coordinates).\ theory Hermitean_Matrices imports Unitary_Matrices begin definition hermitean :: "complex_mat \ bool" where "hermitean A \ mat_adj A = A" lemma hermitean_transpose: shows "hermitean A \ mat_transpose A = mat_cnj A" unfolding hermitean_def by (cases A) (auto simp add: mat_adj_def mat_cnj_def) text \Characterization of 2x2 Hermitean matrices elements. All 2x2 Hermitean matrices are of the form $$ \left( \begin{array}{cc} A & B\\ \overline{B} & D \end{array} \right), $$ for real $A$ and $D$ and complex $B$. \ lemma hermitean_mk_circline [simp]: shows "hermitean (cor A, B, cnj B, cor D)" unfolding hermitean_def mat_adj_def mat_cnj_def by simp lemma hermitean_mk_circline' [simp]: assumes "is_real A" and "is_real D" shows "hermitean (A, B, cnj B, D)" using assms eq_cnj_iff_real unfolding hermitean_def mat_adj_def mat_cnj_def by force lemma hermitean_elems: assumes "hermitean (A, B, C, D)" shows "is_real A" and "is_real D" and "B = cnj C" and "cnj B = C" using assms eq_cnj_iff_real[of A] eq_cnj_iff_real[of D] by (auto simp add: hermitean_def mat_adj_def mat_cnj_def) text \Operations that preserve the Hermitean property\ lemma hermitean_mat_cnj: shows "hermitean H \ hermitean (mat_cnj H)" by (cases H) (auto simp add: hermitean_def mat_adj_def mat_cnj_def) lemma hermitean_mult_real: assumes "hermitean H" shows "hermitean ((cor k) *\<^sub>s\<^sub>m H)" using assms unfolding hermitean_def by simp lemma hermitean_congruence: assumes "hermitean H" shows "hermitean (congruence M H)" using assms unfolding hermitean_def by (auto simp add: mult_mm_assoc) text \Identity matrix is Hermitean\ lemma hermitean_eye [simp]: shows "hermitean eye" by (auto simp add: hermitean_def mat_adj_def mat_cnj_def) lemma hermitean_eye' [simp]: shows "hermitean (1, 0, 0, 1)" by (auto simp add: hermitean_def mat_adj_def mat_cnj_def) text \Unit circle matrix is Hermitean\ lemma hermitean_unit_circle [simp]: shows "hermitean (1, 0, 0, -1)" by (auto simp add: hermitean_def mat_adj_def mat_cnj_def) text \Hermitean matrices have real determinant\ lemma mat_det_hermitean_real: assumes "hermitean A" shows "is_real (mat_det A)" using assms unfolding hermitean_def by (metis eq_cnj_iff_real mat_det_adj) text \Zero matrix is the only Hermitean matrix with both determinant and trace equal to zero\ lemma hermitean_det_zero_trace_zero: assumes "mat_det A = 0" and "mat_trace A = (0::complex)" and "hermitean A" shows "A = mat_zero" using assms proof- { fix a d c assume "a * d = cnj c * c" "a + d = 0" "cnj a = a" from \a + d = 0\ have "d = -a" by (metis add_eq_0_iff) hence "- (cor (Re a))\<^sup>2 = (cor (cmod c))\<^sup>2" using \cnj a = a\ eq_cnj_iff_real[of a] using \a*d = cnj c * c\ using complex_mult_cnj_cmod[of "cnj c"] by (simp add: power2_eq_square) hence "- (Re a)\<^sup>2 \ 0" using zero_le_power2[of "cmod c"] - by (metis Re_complex_of_real cor_squared of_real_minus) + by (metis Re_complex_of_real of_real_minus of_real_power) hence "a = 0" using zero_le_power2[of "Re a"] using \cnj a = a\ eq_cnj_iff_real[of a] by (simp add: complex_eq_if_Re_eq) } note * = this obtain a b c d where "A = (a, b, c, d)" by (cases A) auto thus ?thesis using *[of a d c] *[of d a c] using assms \A = (a, b, c, d)\ by (auto simp add: hermitean_def mat_adj_def mat_cnj_def) qed (* ---------------------------------------------------------------------------- *) subsubsection \Bilinear and quadratic forms with Hermitean matrices\ (* ---------------------------------------------------------------------------- *) text \A Hermitean matrix $(A, B, \overline{B}, D)$, for real $A$ and $D$, gives rise to bilinear form $A\cdot \overline{v_{11}} \cdot v_{21}+\overline{B} \cdot \overline{v_{12}} \cdot v_{21} + B \cdot \overline{v_{11}} \cdot v_{22}+D\cdot \overline{v_{12}}\cdot v_{22}$ (acting on vectors $(v_{11}, v_{12})$ and $(v_{21}, v_{22})$) and to the quadratic form $A \cdot \overline{v_1} \cdot v_1+\overline{B}\cdot \overline{v_2}\cdot v_1 + B\cdot \overline{v_1}\cdot v_2 + D\cdot \overline{v_2} \cdot v_2$ (acting on the vector $(v_1, v_2)$).\ lemma bilinear_form_hermitean_commute: assumes "hermitean H" shows "bilinear_form v1 v2 H = cnj (bilinear_form v2 v1 H)" proof- have "v2 *\<^sub>v\<^sub>m mat_cnj H *\<^sub>v\<^sub>v vec_cnj v1 = vec_cnj v1 *\<^sub>v\<^sub>v (mat_adj H *\<^sub>m\<^sub>v v2)" by (subst mult_vv_commute, subst mult_mv_mult_vm, simp add: mat_adj_def mat_transpose_mat_cnj) also have "\ = bilinear_form v1 v2 H" using assms by (simp add: mult_vv_mv hermitean_def) finally show ?thesis by (simp add: cnj_mult_vv vec_cnj_mult_vm) qed lemma quad_form_hermitean_real: assumes "hermitean H" shows "is_real (quad_form z H)" using assms by (subst eq_cnj_iff_real[symmetric]) (simp del: quad_form_def add: hermitean_def) lemma quad_form_vec_cnj_mat_cnj: assumes "hermitean H" shows "quad_form (vec_cnj z) (mat_cnj H) = quad_form z H" using assms using cnj_mult_vv cnj_quad_form hermitean_def vec_cnj_mult_vm by auto (* ---------------------------------------------------------------------------- *) subsubsection \Eigenvalues, eigenvectors and diagonalization of Hermitean matrices\ (* ---------------------------------------------------------------------------- *) text \Hermitean matrices have real eigenvalues\ lemma hermitean_eigenval_real: assumes "hermitean H" and "eigenval k H" shows "is_real k" proof- from assms obtain v where "v \ vec_zero" "H *\<^sub>m\<^sub>v v = k *\<^sub>s\<^sub>v v" unfolding eigenval_def by blast have "k * (v *\<^sub>v\<^sub>v vec_cnj v) = (k *\<^sub>s\<^sub>v v) *\<^sub>v\<^sub>v (vec_cnj v)" by (simp add: mult_vv_scale_sv1) also have "... = (H *\<^sub>m\<^sub>v v) *\<^sub>v\<^sub>v (vec_cnj v)" using \H *\<^sub>m\<^sub>v v = k *\<^sub>s\<^sub>v v\ by simp also have "... = v *\<^sub>v\<^sub>v (mat_transpose H *\<^sub>m\<^sub>v (vec_cnj v))" by (simp add: mult_mv_vv) also have "... = v *\<^sub>v\<^sub>v (vec_cnj (mat_cnj (mat_transpose H) *\<^sub>m\<^sub>v v))" by (simp add: vec_cnj_mult_mv) also have "... = v *\<^sub>v\<^sub>v (vec_cnj (H *\<^sub>m\<^sub>v v))" using \hermitean H\ by (simp add: hermitean_def mat_adj_def) also have "... = v *\<^sub>v\<^sub>v (vec_cnj (k *\<^sub>s\<^sub>v v))" using \H *\<^sub>m\<^sub>v v = k *\<^sub>s\<^sub>v v\ by simp finally have "k * (v *\<^sub>v\<^sub>v vec_cnj v) = cnj k * (v *\<^sub>v\<^sub>v vec_cnj v)" by (simp add: mult_vv_scale_sv2) hence "k = cnj k" using \v \ vec_zero\ using scalsquare_vv_zero[of v] by (simp add: mult_vv_commute) thus ?thesis by (metis eq_cnj_iff_real) qed text \Non-diagonal Hermitean matrices have distinct eigenvalues\ lemma hermitean_distinct_eigenvals: assumes "hermitean H" shows "(\ k\<^sub>1 k\<^sub>2. k\<^sub>1 \ k\<^sub>2 \ eigenval k\<^sub>1 H \ eigenval k\<^sub>2 H) \ mat_diagonal H" proof- obtain A B C D where HH: "H = (A, B, C, D)" by (cases H) auto show ?thesis proof (cases "B = 0") case True thus ?thesis using \hermitean H\ hermitean_elems[of A B C D] HH by auto next case False have "(mat_trace H)\<^sup>2 \ 4 * mat_det H" proof (rule ccontr) have "C = cnj B" "is_real A" "is_real D" using hermitean_elems HH \hermitean H\ by auto assume "\ ?thesis" hence "(A + D)\<^sup>2 = 4*(A*D - B*C)" using HH by auto hence "(A - D)\<^sup>2 = - 4*B*cnj B" using \C = cnj B\ by (auto simp add: power2_eq_square field_simps) hence "(A - D)\<^sup>2 / cor ((cmod B)\<^sup>2) = -4" using \B \ 0\ complex_mult_cnj_cmod[of B] by (auto simp add: field_simps) hence "(Re A - Re D)\<^sup>2 / (cmod B)\<^sup>2 = -4" using \is_real A\ \is_real D\ \B \ 0\ using Re_divide_real[of "cor ((cmod B)\<^sup>2)" "(A - D)\<^sup>2"] by (auto simp add: power2_eq_square) thus False by (metis abs_neg_numeral abs_power2 neg_numeral_neq_numeral power_divide) qed show ?thesis apply (rule disjI1) apply (subst eigen_equation)+ using complex_quadratic_equation_monic_distinct_roots[of "-mat_trace H" "mat_det H"] \(mat_trace H)\<^sup>2 \ 4 * mat_det H\ by auto qed qed text \Eigenvectors corresponding to different eigenvalues of Hermitean matrices are orthogonal\ lemma hermitean_ortho_eigenvecs: assumes "hermitean H" assumes "eigenpair k1 v1 H" and "eigenpair k2 v2 H" and "k1 \ k2" shows "vec_cnj v2 *\<^sub>v\<^sub>v v1 = 0" and "vec_cnj v1 *\<^sub>v\<^sub>v v2 = 0" proof- from assms have "v1 \ vec_zero" "H *\<^sub>m\<^sub>v v1 = k1 *\<^sub>s\<^sub>v v1" "v2 \ vec_zero" "H *\<^sub>m\<^sub>v v2 = k2 *\<^sub>s\<^sub>v v2" unfolding eigenpair_def by auto have real_k: "is_real k1" "is_real k2" using assms using hermitean_eigenval_real[of H k1] using hermitean_eigenval_real[of H k2] unfolding eigenpair_def eigenval_def by blast+ have "vec_cnj (H *\<^sub>m\<^sub>v v2) = vec_cnj (k2 *\<^sub>s\<^sub>v v2)" using \H *\<^sub>m\<^sub>v v2 = k2 *\<^sub>s\<^sub>v v2\ by auto hence "vec_cnj v2 *\<^sub>v\<^sub>m H = k2 *\<^sub>s\<^sub>v vec_cnj v2" using \hermitean H\ real_k eq_cnj_iff_real[of k1] eq_cnj_iff_real[of k2] unfolding hermitean_def by (cases H, cases v2) (auto simp add: mat_adj_def mat_cnj_def vec_cnj_def) have "k2 * (vec_cnj v2 *\<^sub>v\<^sub>v v1) = k1 * (vec_cnj v2 *\<^sub>v\<^sub>v v1)" using \H *\<^sub>m\<^sub>v v1 = k1 *\<^sub>s\<^sub>v v1\ using \vec_cnj v2 *\<^sub>v\<^sub>m H = k2 *\<^sub>s\<^sub>v vec_cnj v2\ by (cases v1, cases v2, cases H) (metis mult_vv_mv mult_vv_scale_sv1 mult_vv_scale_sv2) thus "vec_cnj v2 *\<^sub>v\<^sub>v v1 = 0" using \k1 \ k2\ by simp hence "cnj (vec_cnj v2 *\<^sub>v\<^sub>v v1) = 0" by simp thus "vec_cnj v1 *\<^sub>v\<^sub>v v2 = 0" by (simp add: cnj_mult_vv mult_vv_commute) qed text \Hermitean matrices are diagonizable by unitary matrices. Diagonal entries are real and the sign of the determinant is preserved.\ lemma hermitean_diagonizable: assumes "hermitean H" shows "\ k1 k2 M. mat_det M \ 0 \ unitary M \ congruence M H = (k1, 0, 0, k2) \ is_real k1 \ is_real k2 \ sgn (Re k1 * Re k2) = sgn (Re (mat_det H))" proof- from assms have "(\k\<^sub>1 k\<^sub>2. k\<^sub>1 \ k\<^sub>2 \ eigenval k\<^sub>1 H \ eigenval k\<^sub>2 H) \ mat_diagonal H" using hermitean_distinct_eigenvals[of H] by simp thus ?thesis proof assume "\k\<^sub>1 k\<^sub>2. k\<^sub>1 \ k\<^sub>2 \ eigenval k\<^sub>1 H \ eigenval k\<^sub>2 H" then obtain k1 k2 where "k1 \ k2" "eigenval k1 H" "eigenval k2 H" using hermitean_distinct_eigenvals by blast then obtain v1 v2 where "eigenpair k1 v1 H" "eigenpair k2 v2 H" "v1 \ vec_zero" "v2 \ vec_zero" unfolding eigenval_def eigenpair_def by blast hence *: "vec_cnj v2 *\<^sub>v\<^sub>v v1 = 0" "vec_cnj v1 *\<^sub>v\<^sub>v v2 = 0" using \k1 \ k2\ hermitean_ortho_eigenvecs \hermitean H\ by auto obtain v11 v12 v21 v22 where vv: "v1 = (v11, v12)" "v2 = (v21, v22)" by (cases v1, cases v2) auto let ?nv1' = "vec_cnj v1 *\<^sub>v\<^sub>v v1" and ?nv2' = "vec_cnj v2 *\<^sub>v\<^sub>v v2" let ?nv1 = "cor (sqrt (Re ?nv1'))" let ?nv2 = "cor (sqrt (Re ?nv2'))" have "?nv1' \ 0" "?nv2' \ 0" using \v1 \ vec_zero\ \v2 \ vec_zero\ vv by (simp add: scalsquare_vv_zero)+ moreover have "is_real ?nv1'" "is_real ?nv2'" using vv by (auto simp add: vec_cnj_def) ultimately have "?nv1 \ 0" "?nv2 \ 0" using complex_eq_if_Re_eq by auto have "Re (?nv1') \ 0" "Re (?nv2') \ 0" using vv by (auto simp add: vec_cnj_def) obtain nv1 nv2 where "nv1 = ?nv1" "nv1 \ 0" "nv2 = ?nv2" "nv2 \ 0" using \?nv1 \ 0\ \?nv2 \ 0\ by auto let ?M = "(1/nv1 * v11, 1/nv2 * v21, 1/nv1 * v12, 1/nv2 * v22)" have "is_real k1" "is_real k2" using \eigenval k1 H\ \eigenval k2 H\ \hermitean H\ by (auto simp add: hermitean_eigenval_real) moreover have "mat_det ?M \ 0" proof (rule ccontr) assume "\ ?thesis" hence "v11 * v22 = v12 * v21" using \nv1 \ 0\ \nv2 \ 0\ by (auto simp add: field_simps) hence "\ k. k \ 0 \ v2 = k *\<^sub>s\<^sub>v v1" using vv \v1 \ vec_zero\ \v2 \ vec_zero\ apply auto apply (rule_tac x="v21/v11" in exI, force simp add: field_simps) apply (rule_tac x="v21/v11" in exI, force simp add: field_simps) apply (rule_tac x="v22/v12" in exI, force simp add: field_simps) apply (rule_tac x="v22/v12" in exI, force simp add: field_simps) done thus False using * \vec_cnj v1 *\<^sub>v\<^sub>v v2 = 0\ \vec_cnj v2 *\<^sub>v\<^sub>v v2 \ 0\ vv \?nv1' \ 0\ by (metis mult_vv_scale_sv2 mult_zero_right) qed moreover have "unitary ?M" proof- have **: "cnj nv1 * nv1 = ?nv1'" "cnj nv2 * nv2 = ?nv2'" using \nv1 = ?nv1\ \nv1 \ 0\ \nv2 = ?nv2\ \nv2 \ 0\ \is_real ?nv1'\ \is_real ?nv2'\ using \Re (?nv1') \ 0\ \Re (?nv2') \ 0\ by auto have ***: "cnj nv1 * nv2 \ 0" "cnj nv2 * nv1 \ 0" using vv \nv1 = ?nv1\ \nv1 \ 0\ \nv2 = ?nv2\ \nv2 \ 0\ \is_real ?nv1'\ \is_real ?nv2'\ by auto show ?thesis unfolding unitary_def using vv ** \?nv1' \ 0\ \?nv2' \ 0\ * *** unfolding mat_adj_def mat_cnj_def vec_cnj_def by simp (metis (no_types, lifting) add_divide_distrib divide_eq_0_iff divide_eq_1_iff) qed moreover have "congruence ?M H = (k1, 0, 0, k2)" proof- have "mat_inv ?M *\<^sub>m\<^sub>m H *\<^sub>m\<^sub>m ?M = (k1, 0, 0, k2)" proof- have *: "H *\<^sub>m\<^sub>m ?M = ?M *\<^sub>m\<^sub>m (k1, 0, 0, k2)" using \eigenpair k1 v1 H\ \eigenpair k2 v2 H\ vv \?nv1 \ 0\ \?nv2 \ 0\ unfolding eigenpair_def vec_cnj_def by (cases H) (smt mult_mm.simps vec_map.simps add.right_neutral add_cancel_left_left distrib_left fst_mult_sv mult.commute mult.left_commute mult_mv.simps mult_zero_right prod.sel(1) prod.sel(2) snd_mult_sv) show ?thesis using mult_mm_inv_l[of ?M "(k1, 0, 0, k2)" "H *\<^sub>m\<^sub>m ?M", OF \mat_det ?M \ 0\ *[symmetric], symmetric] by (simp add: mult_mm_assoc) qed moreover have "mat_inv ?M = mat_adj ?M" using \mat_det ?M \ 0\ \unitary ?M\ mult_mm_inv_r[of ?M "mat_adj ?M" eye] by (simp add: unitary_def) ultimately show ?thesis by simp qed moreover have "sgn (Re k1 * Re k2) = sgn (Re (mat_det H))" using \congruence ?M H = (k1, 0, 0, k2)\ \is_real k1\ \is_real k2\ using Re_det_sgn_congruence[of ?M H] \mat_det ?M \ 0\ \hermitean H\ by simp ultimately show ?thesis by (rule_tac x="k1" in exI, rule_tac x="k2" in exI, rule_tac x="?M" in exI) simp next assume "mat_diagonal H" then obtain A D where "H = (A, 0, 0, D)" by (cases H) auto moreover hence "is_real A" "is_real D" using \hermitean H\ hermitean_elems[of A 0 0 D] by auto ultimately show ?thesis by (rule_tac x="A" in exI, rule_tac x="D" in exI, rule_tac x="eye" in exI) (simp add: unitary_def mat_adj_def mat_cnj_def) qed qed end diff --git a/thys/Complex_Geometry/Matrices.thy b/thys/Complex_Geometry/Matrices.thy --- a/thys/Complex_Geometry/Matrices.thy +++ b/thys/Complex_Geometry/Matrices.thy @@ -1,830 +1,830 @@ (* ---------------------------------------------------------------------------- *) subsection \Vectors and Matrices in $\mathbb{C}^2$\ (* ---------------------------------------------------------------------------- *) text \Representing vectors and matrices of arbitrary dimensions pose a challenge in formal theorem proving \cite{harrison05}, but we only need to consider finite dimension spaces $\mathbb{C}^2$ and $\mathbb{R}^3$.\ theory Matrices imports More_Complex Linear_Systems Quadratic begin (* ---------------------------------------------------------------------------- *) subsubsection \Vectors in $\mathbb{C}^2$\ (* ---------------------------------------------------------------------------- *) text \Type of complex vector\ type_synonym complex_vec = "complex \ complex" definition vec_zero :: "complex_vec" where [simp]: "vec_zero = (0, 0)" text \Vector scalar multiplication\ fun mult_sv :: "complex \ complex_vec \ complex_vec" (infixl "*\<^sub>s\<^sub>v" 100) where "k *\<^sub>s\<^sub>v (x, y) = (k*x, k*y)" lemma fst_mult_sv [simp]: shows "fst (k *\<^sub>s\<^sub>v v) = k * fst v" by (cases v) simp lemma snd_mult_sv [simp]: shows "snd (k *\<^sub>s\<^sub>v v) = k * snd v" by (cases v) simp lemma mult_sv_mult_sv [simp]: shows "k1 *\<^sub>s\<^sub>v (k2 *\<^sub>s\<^sub>v v) = (k1*k2) *\<^sub>s\<^sub>v v" by (cases v) simp lemma one_mult_sv [simp]: shows "1 *\<^sub>s\<^sub>v v = v" by (cases v) simp lemma mult_sv_ex_id1 [simp]: shows "\ k::complex. k \ 0 \ k *\<^sub>s\<^sub>v v = v" by (rule_tac x=1 in exI, simp) lemma mult_sv_ex_id2 [simp]: shows "\ k::complex. k \ 0 \ v = k *\<^sub>s\<^sub>v v" by (rule_tac x=1 in exI, simp) text \Scalar product of two vectors\ fun mult_vv :: "complex \ complex \ complex \ complex \ complex" (infixl "*\<^sub>v\<^sub>v" 100) where "(x, y) *\<^sub>v\<^sub>v (a, b) = x*a + y*b" lemma mult_vv_commute: shows "v1 *\<^sub>v\<^sub>v v2 = v2 *\<^sub>v\<^sub>v v1" by (cases v1, cases v2) auto lemma mult_vv_scale_sv1: shows "(k *\<^sub>s\<^sub>v v1) *\<^sub>v\<^sub>v v2 = k * (v1 *\<^sub>v\<^sub>v v2)" by (cases v1, cases v2) (auto simp add: field_simps) lemma mult_vv_scale_sv2: shows "v1 *\<^sub>v\<^sub>v (k *\<^sub>s\<^sub>v v2) = k * (v1 *\<^sub>v\<^sub>v v2)" by (cases v1, cases v2) (auto simp add: field_simps) text \Conjugate vector\ fun vec_map where "vec_map f (x, y) = (f x, f y)" definition vec_cnj where "vec_cnj = vec_map cnj" lemma vec_cnj_vec_cnj [simp]: shows "vec_cnj (vec_cnj v) = v" by (cases v) (simp add: vec_cnj_def) lemma cnj_mult_vv: shows "cnj (v1 *\<^sub>v\<^sub>v v2) = (vec_cnj v1) *\<^sub>v\<^sub>v (vec_cnj v2)" by (cases v1, cases v2) (simp add: vec_cnj_def) lemma vec_cnj_sv [simp]: shows "vec_cnj (k *\<^sub>s\<^sub>v A) = cnj k *\<^sub>s\<^sub>v vec_cnj A" by (cases A) (auto simp add: vec_cnj_def) lemma scalsquare_vv_zero: shows "(vec_cnj v) *\<^sub>v\<^sub>v v = 0 \ v = vec_zero" apply (cases v) apply (auto simp add: vec_cnj_def field_simps complex_mult_cnj_cmod power2_eq_square) - apply (simp only: cor_add[symmetric] cor_mult[symmetric] of_real_eq_0_iff, simp)+ + apply (metis (no_types) norm_eq_zero of_real_0 of_real_add of_real_eq_iff of_real_mult sum_squares_eq_zero_iff)+ done (* ---------------------------------------------------------------------------- *) subsubsection \Matrices in $\mathbb{C}^2$\ (* ---------------------------------------------------------------------------- *) text \Type of complex matrices\ type_synonym complex_mat = "complex \ complex \ complex \ complex" text \Matrix scalar multiplication\ fun mult_sm :: "complex \ complex_mat \ complex_mat" (infixl "*\<^sub>s\<^sub>m" 100) where "k *\<^sub>s\<^sub>m (a, b, c, d) = (k*a, k*b, k*c, k*d)" lemma mult_sm_distribution [simp]: shows "k1 *\<^sub>s\<^sub>m (k2 *\<^sub>s\<^sub>m A) = (k1*k2) *\<^sub>s\<^sub>m A" by (cases A) auto lemma mult_sm_neutral [simp]: shows "1 *\<^sub>s\<^sub>m A = A" by (cases A) auto lemma mult_sm_inv_l: assumes "k \ 0" and "k *\<^sub>s\<^sub>m A = B" shows "A = (1/k) *\<^sub>s\<^sub>m B" using assms by auto lemma mult_sm_ex_id1 [simp]: shows "\ k::complex. k \ 0 \ k *\<^sub>s\<^sub>m M = M" by (rule_tac x=1 in exI, simp) lemma mult_sm_ex_id2 [simp]: shows "\ k::complex. k \ 0 \ M = k *\<^sub>s\<^sub>m M" by (rule_tac x=1 in exI, simp) text \Matrix addition and subtraction\ definition mat_zero :: "complex_mat" where [simp]: "mat_zero = (0, 0, 0, 0)" fun mat_plus :: "complex_mat \ complex_mat \ complex_mat" (infixl "+\<^sub>m\<^sub>m" 100) where "mat_plus (a1, b1, c1, d1) (a2, b2, c2, d2) = (a1+a2, b1+b2, c1+c2, d1+d2)" fun mat_minus :: "complex_mat \ complex_mat \ complex_mat" (infixl "-\<^sub>m\<^sub>m" 100) where "mat_minus (a1, b1, c1, d1) (a2, b2, c2, d2) = (a1-a2, b1-b2, c1-c2, d1-d2)" fun mat_uminus :: "complex_mat \ complex_mat" where "mat_uminus (a, b, c, d) = (-a, -b, -c, -d)" lemma nonzero_mult_real: assumes "A \ mat_zero" and "k \ 0" shows "k *\<^sub>s\<^sub>m A \ mat_zero" using assms by (cases A) simp text \Matrix multiplication.\ fun mult_mm :: "complex_mat \ complex_mat \ complex_mat" (infixl "*\<^sub>m\<^sub>m" 100) where "(a1, b1, c1, d1) *\<^sub>m\<^sub>m (a2, b2, c2, d2) = (a1*a2 + b1*c2, a1*b2 + b1*d2, c1*a2+d1*c2, c1*b2+d1*d2)" lemma mult_mm_assoc: shows "A *\<^sub>m\<^sub>m (B *\<^sub>m\<^sub>m C) = (A *\<^sub>m\<^sub>m B) *\<^sub>m\<^sub>m C" by (cases A, cases B, cases C) (auto simp add: field_simps) lemma mult_assoc_5: shows "A *\<^sub>m\<^sub>m (B *\<^sub>m\<^sub>m C *\<^sub>m\<^sub>m D) *\<^sub>m\<^sub>m E = (A *\<^sub>m\<^sub>m B) *\<^sub>m\<^sub>m C *\<^sub>m\<^sub>m (D *\<^sub>m\<^sub>m E)" by (simp only: mult_mm_assoc) lemma mat_zero_r [simp]: shows "A *\<^sub>m\<^sub>m mat_zero = mat_zero" by (cases A) simp lemma mat_zero_l [simp]: shows "mat_zero *\<^sub>m\<^sub>m A = mat_zero" by (cases A) simp definition eye :: "complex_mat" where [simp]: "eye = (1, 0, 0, 1)" lemma mat_eye_l: shows "eye *\<^sub>m\<^sub>m A = A" by (cases A) auto lemma mat_eye_r: shows "A *\<^sub>m\<^sub>m eye = A" by (cases A) auto lemma mult_mm_sm [simp]: shows "A *\<^sub>m\<^sub>m (k *\<^sub>s\<^sub>m B) = k *\<^sub>s\<^sub>m (A *\<^sub>m\<^sub>m B)" by (cases A, cases B) (simp add: field_simps) lemma mult_sm_mm [simp]: shows "(k *\<^sub>s\<^sub>m A) *\<^sub>m\<^sub>m B = k *\<^sub>s\<^sub>m (A *\<^sub>m\<^sub>m B)" by (cases A, cases B) (simp add: field_simps) lemma mult_sm_eye_mm [simp]: shows "k *\<^sub>s\<^sub>m eye *\<^sub>m\<^sub>m A = k *\<^sub>s\<^sub>m A" by (cases A) simp text \Matrix determinant\ fun mat_det where "mat_det (a, b, c, d) = a*d - b*c" lemma mat_det_mult [simp]: shows "mat_det (A *\<^sub>m\<^sub>m B) = mat_det A * mat_det B" by (cases A, cases B) (auto simp add: field_simps) lemma mat_det_mult_sm [simp]: shows "mat_det (k *\<^sub>s\<^sub>m A) = (k*k) * mat_det A" by (cases A) (auto simp add: field_simps) text \Matrix inverse\ fun mat_inv :: "complex_mat \ complex_mat" where "mat_inv (a, b, c, d) = (1/(a*d - b*c)) *\<^sub>s\<^sub>m (d, -b, -c, a)" lemma mat_inv_r: assumes "mat_det A \ 0" shows "A *\<^sub>m\<^sub>m (mat_inv A) = eye" using assms proof (cases A, auto simp add: field_simps) fix a b c d :: complex assume "a * (a * (d * d)) + b * (b * (c * c)) = a * (b * (c * (d * 2)))" hence "(a*d - b*c)*(a*d - b*c) = 0" by (auto simp add: field_simps) hence *: "a*d - b*c = 0" by auto assume "a*d \ b*c" with * show False by auto qed lemma mat_inv_l: assumes "mat_det A \ 0" shows "(mat_inv A) *\<^sub>m\<^sub>m A = eye" using assms proof (cases A, auto simp add: field_simps) fix a b c d :: complex assume "a * (a * (d * d)) + b * (b * (c * c)) = a * (b * (c * (d * 2)))" hence "(a*d - b*c)*(a*d - b*c) = 0" by (auto simp add: field_simps) hence *: "a*d - b*c = 0" by auto assume "a*d \ b*c" with * show False by auto qed lemma mat_det_inv: assumes "mat_det A \ 0" shows "mat_det (mat_inv A) = 1 / mat_det A" proof- have "mat_det eye = mat_det A * mat_det (mat_inv A)" using mat_inv_l[OF assms, symmetric] by simp thus ?thesis using assms by (simp add: field_simps) qed lemma mult_mm_inv_l: assumes "mat_det A \ 0" and "A *\<^sub>m\<^sub>m B = C" shows "B = mat_inv A *\<^sub>m\<^sub>m C" using assms mat_eye_l[of B] by (auto simp add: mult_mm_assoc mat_inv_l) lemma mult_mm_inv_r: assumes "mat_det B \ 0" and "A *\<^sub>m\<^sub>m B = C" shows "A = C *\<^sub>m\<^sub>m mat_inv B" using assms mat_eye_r[of A] by (auto simp add: mult_mm_assoc[symmetric] mat_inv_r) lemma mult_mm_non_zero_l: assumes "mat_det A \ 0" and "B \ mat_zero" shows "A *\<^sub>m\<^sub>m B \ mat_zero" using assms mat_zero_r using mult_mm_inv_l[OF assms(1), of B mat_zero] by auto lemma mat_inv_mult_mm: assumes "mat_det A \ 0" and "mat_det B \ 0" shows "mat_inv (A *\<^sub>m\<^sub>m B) = mat_inv B *\<^sub>m\<^sub>m mat_inv A" using assms proof- have "(A *\<^sub>m\<^sub>m B) *\<^sub>m\<^sub>m (mat_inv B *\<^sub>m\<^sub>m mat_inv A) = eye" using assms by (metis mat_inv_r mult_mm_assoc mult_mm_inv_r) thus ?thesis using mult_mm_inv_l[of "A *\<^sub>m\<^sub>m B" "mat_inv B *\<^sub>m\<^sub>m mat_inv A" eye] assms mat_eye_r by simp qed lemma mult_mm_cancel_l: assumes "mat_det M \ 0" "M *\<^sub>m\<^sub>m A = M *\<^sub>m\<^sub>m B" shows "A = B" using assms by (metis mult_mm_inv_l) lemma mult_mm_cancel_r: assumes "mat_det M \ 0" "A *\<^sub>m\<^sub>m M = B *\<^sub>m\<^sub>m M" shows "A = B" using assms by (metis mult_mm_inv_r) lemma mult_mm_non_zero_r: assumes "A \ mat_zero" and "mat_det B \ 0" shows "A *\<^sub>m\<^sub>m B \ mat_zero" using assms mat_zero_l using mult_mm_inv_r[OF assms(2), of A mat_zero] by auto lemma mat_inv_mult_sm: assumes "k \ 0" shows "mat_inv (k *\<^sub>s\<^sub>m A) = (1 / k) *\<^sub>s\<^sub>m mat_inv A" proof- obtain a b c d where "A = (a, b, c, d)" by (cases A) auto thus ?thesis using assms by auto (subst mult.assoc[of k a "k*d"], subst mult.assoc[of k b "k*c"], subst right_diff_distrib[of k "a*(k*d)" "b*(k*c)", symmetric], simp, simp add: field_simps)+ qed lemma mat_inv_inv [simp]: assumes "mat_det M \ 0" shows "mat_inv (mat_inv M) = M" proof- have "mat_inv M *\<^sub>m\<^sub>m M = eye" using mat_inv_l[OF assms] by simp thus ?thesis using assms mat_det_inv[of M] using mult_mm_inv_l[of "mat_inv M" M eye] mat_eye_r by (auto simp del: eye_def) qed text \Matrix transpose\ fun mat_transpose where "mat_transpose (a, b, c, d) = (a, c, b, d)" lemma mat_t_mat_t [simp]: shows "mat_transpose (mat_transpose A) = A" by (cases A) auto lemma mat_t_mult_sm [simp]: shows "mat_transpose (k *\<^sub>s\<^sub>m A) = k *\<^sub>s\<^sub>m (mat_transpose A)" by (cases A) simp lemma mat_t_mult_mm [simp]: shows "mat_transpose (A *\<^sub>m\<^sub>m B) = mat_transpose B *\<^sub>m\<^sub>m mat_transpose A" by (cases A, cases B) auto lemma mat_inv_transpose: shows "mat_transpose (mat_inv M) = mat_inv (mat_transpose M)" by (cases M) auto lemma mat_det_transpose [simp]: fixes M :: "complex_mat" shows "mat_det (mat_transpose M) = mat_det M" by (cases M) auto text \Diagonal matrices definition\ fun mat_diagonal where "mat_diagonal (A, B, C, D) = (B = 0 \ C = 0)" text \Matrix conjugate\ fun mat_map where "mat_map f (a, b, c, d) = (f a, f b, f c, f d)" definition mat_cnj where "mat_cnj = mat_map cnj" lemma mat_cnj_cnj [simp]: shows "mat_cnj (mat_cnj A) = A" unfolding mat_cnj_def by (cases A) auto lemma mat_cnj_sm [simp]: shows "mat_cnj (k *\<^sub>s\<^sub>m A) = cnj k *\<^sub>s\<^sub>m (mat_cnj A)" by (cases A) (simp add: mat_cnj_def) lemma mat_det_cnj [simp]: shows "mat_det (mat_cnj A) = cnj (mat_det A)" by (cases A) (simp add: mat_cnj_def) lemma nonzero_mat_cnj: shows "mat_cnj A = mat_zero \ A = mat_zero" by (cases A) (auto simp add: mat_cnj_def) lemma mat_inv_cnj: shows "mat_cnj (mat_inv M) = mat_inv (mat_cnj M)" unfolding mat_cnj_def by (cases M) auto text \Matrix adjoint - the conjugate traspose matrix ($A^* = \overline{A^t}$)\ definition mat_adj where "mat_adj A = mat_cnj (mat_transpose A)" lemma mat_adj_mult_mm [simp]: shows "mat_adj (A *\<^sub>m\<^sub>m B) = mat_adj B *\<^sub>m\<^sub>m mat_adj A" by (cases A, cases B) (auto simp add: mat_adj_def mat_cnj_def) lemma mat_adj_mult_sm [simp]: shows "mat_adj (k *\<^sub>s\<^sub>m A) = cnj k *\<^sub>s\<^sub>m mat_adj A" by (cases A) (auto simp add: mat_adj_def mat_cnj_def) lemma mat_det_adj: shows "mat_det (mat_adj A) = cnj (mat_det A)" by (cases A) (auto simp add: mat_adj_def mat_cnj_def) lemma mat_adj_inv: assumes "mat_det M \ 0" shows "mat_adj (mat_inv M) = mat_inv (mat_adj M)" by (cases M) (auto simp add: mat_adj_def mat_cnj_def) lemma mat_transpose_mat_cnj: shows "mat_transpose (mat_cnj A) = mat_adj A" by (cases A) (auto simp add: mat_adj_def mat_cnj_def) lemma mat_adj_adj [simp]: shows "mat_adj (mat_adj A) = A" unfolding mat_adj_def by (subst mat_transpose_mat_cnj) (simp add: mat_adj_def) lemma mat_adj_eye [simp]: shows "mat_adj eye = eye" by (auto simp add: mat_adj_def mat_cnj_def) text \Matrix trace\ fun mat_trace where "mat_trace (a, b, c, d) = a + d" text \Multiplication of matrix and a vector\ fun mult_mv :: "complex_mat \ complex_vec \ complex_vec" (infixl "*\<^sub>m\<^sub>v" 100) where "(a, b, c, d) *\<^sub>m\<^sub>v (x, y) = (x*a + y*b, x*c + y*d)" fun mult_vm :: "complex_vec \ complex_mat \ complex_vec" (infixl "*\<^sub>v\<^sub>m" 100) where "(x, y) *\<^sub>v\<^sub>m (a, b, c, d) = (x*a + y*c, x*b + y*d)" lemma eye_mv_l [simp]: shows "eye *\<^sub>m\<^sub>v v = v" by (cases v) simp lemma mult_mv_mv [simp]: shows "B *\<^sub>m\<^sub>v (A *\<^sub>m\<^sub>v v) = (B *\<^sub>m\<^sub>m A) *\<^sub>m\<^sub>v v" by (cases v, cases A, cases B) (auto simp add: field_simps) lemma mult_vm_vm [simp]: shows "(v *\<^sub>v\<^sub>m A) *\<^sub>v\<^sub>m B = v *\<^sub>v\<^sub>m (A *\<^sub>m\<^sub>m B)" by (cases v, cases A, cases B) (auto simp add: field_simps) lemma mult_mv_inv: assumes "x = A *\<^sub>m\<^sub>v y" and "mat_det A \ 0" shows "y = (mat_inv A) *\<^sub>m\<^sub>v x" using assms by (cases y) (simp add: mat_inv_l) lemma mult_vm_inv: assumes "x = y *\<^sub>v\<^sub>m A" and "mat_det A \ 0" shows "y = x *\<^sub>v\<^sub>m (mat_inv A) " using assms by (cases y) (simp add: mat_inv_r) lemma mult_mv_cancel_l: assumes "mat_det A \ 0" and "A *\<^sub>m\<^sub>v v = A *\<^sub>m\<^sub>v v'" shows "v = v'" using assms using mult_mv_inv by blast lemma mult_vm_cancel_r: assumes "mat_det A \ 0" and "v *\<^sub>v\<^sub>m A = v' *\<^sub>v\<^sub>m A" shows "v = v'" using assms using mult_vm_inv by blast lemma vec_zero_l [simp]: shows "A *\<^sub>m\<^sub>v vec_zero = vec_zero" by (cases A) simp lemma vec_zero_r [simp]: shows "vec_zero *\<^sub>v\<^sub>m A = vec_zero" by (cases A) simp lemma mult_mv_nonzero: assumes "v \ vec_zero" and "mat_det A \ 0" shows "A *\<^sub>m\<^sub>v v \ vec_zero" apply (rule ccontr) using assms mult_mv_inv[of vec_zero A v] mat_inv_l vec_zero_l by auto lemma mult_vm_nonzero: assumes "v \ vec_zero" and "mat_det A \ 0" shows "v *\<^sub>v\<^sub>m A \ vec_zero" apply (rule ccontr) using assms mult_vm_inv[of vec_zero v A] mat_inv_r vec_zero_r by auto lemma mult_sv_mv: shows "k *\<^sub>s\<^sub>v (A *\<^sub>m\<^sub>v v) = (A *\<^sub>m\<^sub>v (k *\<^sub>s\<^sub>v v))" by (cases A, cases v) (simp add: field_simps) lemma mult_mv_mult_vm: shows "A *\<^sub>m\<^sub>v x = x *\<^sub>v\<^sub>m (mat_transpose A)" by (cases A, cases x) auto lemma mult_mv_vv: shows "A *\<^sub>m\<^sub>v v1 *\<^sub>v\<^sub>v v2 = v1 *\<^sub>v\<^sub>v (mat_transpose A *\<^sub>m\<^sub>v v2)" by (cases v1, cases v2, cases A) (auto simp add: field_simps) lemma mult_vv_mv: shows "x *\<^sub>v\<^sub>v (A *\<^sub>m\<^sub>v y) = (x *\<^sub>v\<^sub>m A) *\<^sub>v\<^sub>v y" by (cases x, cases y, cases A) (auto simp add: field_simps) lemma vec_cnj_mult_mv: shows "vec_cnj (A *\<^sub>m\<^sub>v x) = (mat_cnj A) *\<^sub>m\<^sub>v (vec_cnj x)" by (cases A, cases x) (auto simp add: vec_cnj_def mat_cnj_def) lemma vec_cnj_mult_vm: shows "vec_cnj (v *\<^sub>v\<^sub>m A) = vec_cnj v *\<^sub>v\<^sub>m mat_cnj A" unfolding vec_cnj_def mat_cnj_def by (cases A, cases v, auto) (* ---------------------------------------------------------------------------- *) subsubsection \Eigenvalues and eigenvectors\ (* ---------------------------------------------------------------------------- *) definition eigenpair where [simp]: "eigenpair k v H \ v \ vec_zero \ H *\<^sub>m\<^sub>v v = k *\<^sub>s\<^sub>v v" definition eigenval where [simp]: "eigenval k H \ (\ v. v \ vec_zero \ H *\<^sub>m\<^sub>v v = k *\<^sub>s\<^sub>v v)" lemma eigen_equation: shows "eigenval k H \ k\<^sup>2 - mat_trace H * k + mat_det H = 0" (is "?lhs \ ?rhs") proof- obtain A B C D where HH: "H = (A, B, C, D)" by (cases H) auto show ?thesis proof assume ?lhs then obtain v where "v \ vec_zero" "H *\<^sub>m\<^sub>v v = k *\<^sub>s\<^sub>v v" unfolding eigenval_def by blast obtain v1 v2 where vv: "v = (v1, v2)" by (cases v) auto from \H *\<^sub>m\<^sub>v v = k *\<^sub>s\<^sub>v v\ have "(H -\<^sub>m\<^sub>m (k *\<^sub>s\<^sub>m eye)) *\<^sub>m\<^sub>v v = vec_zero" using HH vv by (auto simp add: field_simps) hence "mat_det (H -\<^sub>m\<^sub>m (k *\<^sub>s\<^sub>m eye)) = 0" using \v \ vec_zero\ vv HH using regular_homogenous_system[of "A - k" B C "D - k" v1 v2] unfolding det2_def by (auto simp add: field_simps) thus ?rhs using HH by (auto simp add: power2_eq_square field_simps) next assume ?rhs hence *: "mat_det (H -\<^sub>m\<^sub>m (k *\<^sub>s\<^sub>m eye)) = 0" using HH by (auto simp add: field_simps power2_eq_square) show ?lhs proof (cases "H -\<^sub>m\<^sub>m (k *\<^sub>s\<^sub>m eye) = mat_zero") case True thus ?thesis using HH by (auto) (rule_tac x=1 in exI, simp) next case False hence "(A - k \ 0 \ B \ 0) \ (D - k \ 0 \ C \ 0)" using HH by auto thus ?thesis proof assume "A - k \ 0 \ B \ 0" hence "C * B + (D - k) * (k - A) = 0" using * singular_system[of "A-k" "D-k" B C "(0, 0)" 0 0 "(B, k-A)"] HH by (auto simp add: field_simps) hence "(B, k-A) \ vec_zero" "(H -\<^sub>m\<^sub>m (k *\<^sub>s\<^sub>m eye)) *\<^sub>m\<^sub>v (B, k-A) = vec_zero" using HH \A - k \ 0 \ B \ 0\ by (auto simp add: field_simps) then obtain v where "v \ vec_zero \ (H -\<^sub>m\<^sub>m (k *\<^sub>s\<^sub>m eye)) *\<^sub>m\<^sub>v v = vec_zero" by blast thus ?thesis using HH unfolding eigenval_def by (rule_tac x="v" in exI) (case_tac v, simp add: field_simps) next assume "D - k \ 0 \ C \ 0" hence "C * B + (D - k) * (k - A) = 0" using * singular_system[of "D-k" "A-k" C B "(0, 0)" 0 0 "(C, k-D)"] HH by (auto simp add: field_simps) hence "(k-D, C) \ vec_zero" "(H -\<^sub>m\<^sub>m (k *\<^sub>s\<^sub>m eye)) *\<^sub>m\<^sub>v (k-D, C) = vec_zero" using HH \D - k \ 0 \ C \ 0\ by (auto simp add: field_simps) then obtain v where "v \ vec_zero \ (H -\<^sub>m\<^sub>m (k *\<^sub>s\<^sub>m eye)) *\<^sub>m\<^sub>v v = vec_zero" by blast thus ?thesis using HH unfolding eigenval_def by (rule_tac x="v" in exI) (case_tac v, simp add: field_simps) qed qed qed qed (* ---------------------------------------------------------------------------- *) subsubsection \Bilinear and Quadratic forms, Congruence, and Similarity\ (* ---------------------------------------------------------------------------- *) text \Bilinear forms\ definition bilinear_form where [simp]: "bilinear_form v1 v2 H = (vec_cnj v1) *\<^sub>v\<^sub>m H *\<^sub>v\<^sub>v v2" lemma bilinear_form_scale_m: shows "bilinear_form v1 v2 (k *\<^sub>s\<^sub>m H) = k * bilinear_form v1 v2 H" by (cases v1, cases v2, cases H) (simp add: vec_cnj_def field_simps) lemma bilinear_form_scale_v1: shows "bilinear_form (k *\<^sub>s\<^sub>v v1) v2 H = cnj k * bilinear_form v1 v2 H" by (cases v1, cases v2, cases H) (simp add: vec_cnj_def field_simps) lemma bilinear_form_scale_v2: shows "bilinear_form v1 (k *\<^sub>s\<^sub>v v2) H = k * bilinear_form v1 v2 H" by (cases v1, cases v2, cases H) (simp add: vec_cnj_def field_simps) text \Quadratic forms\ definition quad_form where [simp]: "quad_form v H = (vec_cnj v) *\<^sub>v\<^sub>m H *\<^sub>v\<^sub>v v" lemma quad_form_bilinear_form: shows "quad_form v H = bilinear_form v v H" by simp lemma quad_form_scale_v: shows "quad_form (k *\<^sub>s\<^sub>v v) H = cor ((cmod k)\<^sup>2) * quad_form v H" using bilinear_form_scale_v1 bilinear_form_scale_v2 by (simp add: complex_mult_cnj_cmod field_simps) lemma quad_form_scale_m: shows "quad_form v (k *\<^sub>s\<^sub>m H) = k * quad_form v H" using bilinear_form_scale_m by simp lemma cnj_quad_form [simp]: shows "cnj (quad_form z H) = quad_form z (mat_adj H)" by (cases H, cases z) (auto simp add: mat_adj_def mat_cnj_def vec_cnj_def field_simps) text \Matrix congruence\ text \Two matrices are congruent iff they represent the same quadratic form with respect to different bases (for example if one circline can be transformed to another by a Möbius trasformation).\ definition congruence where [simp]: "congruence M H \ mat_adj M *\<^sub>m\<^sub>m H *\<^sub>m\<^sub>m M" lemma congruence_nonzero: assumes "H \ mat_zero" and "mat_det M \ 0" shows "congruence M H \ mat_zero" using assms unfolding congruence_def by (subst mult_mm_non_zero_r, subst mult_mm_non_zero_l) (auto simp add: mat_det_adj) lemma congruence_congruence: shows "congruence M1 (congruence M2 H) = congruence (M2 *\<^sub>m\<^sub>m M1) H" unfolding congruence_def apply (subst mult_mm_assoc) apply (subst mult_mm_assoc) apply (subst mat_adj_mult_mm) apply (subst mult_mm_assoc) by simp lemma congruence_eye [simp]: shows "congruence eye H = H" by (cases H) (simp add: mat_adj_def mat_cnj_def) lemma congruence_congruence_inv [simp]: assumes "mat_det M \ 0" shows "congruence M (congruence (mat_inv M) H) = H" using assms congruence_congruence[of M "mat_inv M" H] using mat_inv_l[of M] mat_eye_l mat_eye_r unfolding congruence_def by (simp del: eye_def) lemma congruence_inv: assumes "mat_det M \ 0" and "congruence M H = H'" shows "congruence (mat_inv M) H' = H" using assms using \mat_det M \ 0\ mult_mm_inv_l[of "mat_adj M" "H *\<^sub>m\<^sub>m M" "H'"] using mult_mm_inv_r[of M "H" "mat_inv (mat_adj M) *\<^sub>m\<^sub>m H'"] by (simp add: mat_det_adj mult_mm_assoc mat_adj_inv) lemma congruence_scale_m [simp]: shows "congruence M (k *\<^sub>s\<^sub>m H) = k *\<^sub>s\<^sub>m (congruence M H)" by (cases M, cases H) (auto simp add: mat_adj_def mat_cnj_def field_simps) lemma inj_congruence: assumes "mat_det M \ 0" and "congruence M H = congruence M H'" shows "H = H'" proof- have "H *\<^sub>m\<^sub>m M = H' *\<^sub>m\<^sub>m M " using assms using mult_mm_cancel_l[of "mat_adj M" "H *\<^sub>m\<^sub>m M" "H' *\<^sub>m\<^sub>m M"] by (simp add: mat_det_adj mult_mm_assoc) thus ?thesis using assms using mult_mm_cancel_r[of "M" "H" "H'"] by simp qed lemma mat_det_congruence [simp]: "mat_det (congruence M H) = (cor ((cmod (mat_det M))\<^sup>2)) * mat_det H" using complex_mult_cnj_cmod[of "mat_det M"] by (auto simp add: mat_det_adj field_simps) lemma det_sgn_congruence [simp]: assumes "mat_det M \ 0" shows "sgn (mat_det (congruence M H)) = sgn (mat_det H)" using assms by (subst mat_det_congruence, auto simp add: sgn_mult power2_eq_square) (simp add: sgn_of_real) lemma Re_det_sgn_congruence [simp]: assumes "mat_det M \ 0" shows "sgn (Re (mat_det (congruence M H))) = sgn (Re (mat_det H))" proof- have *: "Re (mat_det (congruence M H)) = (cmod (mat_det M))\<^sup>2 * Re (mat_det H)" by (subst mat_det_congruence, subst Re_mult_real, rule Im_complex_of_real) (subst Re_complex_of_real, simp) show ?thesis using assms by (subst *) (auto simp add: sgn_mult) qed text \Transforming a matrix $H$ by a regular matrix $M$ preserves its bilinear and quadratic forms.\ lemma bilinear_form_congruence [simp]: assumes "mat_det M \ 0" shows "bilinear_form (M *\<^sub>m\<^sub>v v1) (M *\<^sub>m\<^sub>v v2) (congruence (mat_inv M) H) = bilinear_form v1 v2 H" proof- have "mat_det (mat_adj M) \ 0" using assms by (simp add: mat_det_adj) show ?thesis unfolding bilinear_form_def congruence_def apply (subst mult_mv_mult_vm) apply (subst vec_cnj_mult_vm) apply (subst mat_adj_def[symmetric]) apply (subst mult_vm_vm) apply (subst mult_vv_mv) apply (subst mult_vm_vm) apply (subst mat_adj_inv[OF \mat_det M \ 0\]) apply (subst mult_assoc_5) apply (subst mat_inv_r[OF \mat_det (mat_adj M) \ 0\]) apply (subst mat_inv_l[OF \mat_det M \ 0\]) apply (subst mat_eye_l, subst mat_eye_r) by simp qed lemma quad_form_congruence [simp]: assumes "mat_det M \ 0" shows "quad_form (M *\<^sub>m\<^sub>v z) (congruence (mat_inv M) H) = quad_form z H" using bilinear_form_congruence[OF assms] by simp text \Similar matrices\ text \Two matrices are similar iff they represent the same linear operator with respect to (possibly) different bases (e.g., if they represent the same Möbius transformation after changing the coordinate system)\ definition similarity where "similarity A M = mat_inv A *\<^sub>m\<^sub>m M *\<^sub>m\<^sub>m A" lemma mat_det_similarity [simp]: assumes "mat_det A \ 0" shows "mat_det (similarity A M) = mat_det M" using assms unfolding similarity_def by (simp add: mat_det_inv) lemma mat_trace_similarity [simp]: assumes "mat_det A \ 0" shows "mat_trace (similarity A M) = mat_trace M" proof- obtain a b c d where AA: "A = (a, b, c, d)" by (cases A) auto obtain mA mB mC mD where MM: "M = (mA, mB, mC, mD)" by (cases M) auto have "mA * (a * d) / (a * d - b * c) + mD * (a * d) / (a * d - b * c) = mA + mD + mA * (b * c) / (a * d - b * c) + mD * (b * c) / (a * d - b * c)" using assms AA by (simp add: field_simps) thus ?thesis using AA MM by (simp add: field_simps similarity_def) qed lemma similarity_eye [simp]: shows "similarity eye M = M" unfolding similarity_def using mat_eye_l mat_eye_r by auto lemma similarity_eye' [simp]: shows "similarity (1, 0, 0, 1) M = M" unfolding eye_def[symmetric] by (simp del: eye_def) lemma similarity_comp [simp]: assumes "mat_det A1 \ 0" and "mat_det A2 \ 0" shows "similarity A1 (similarity A2 M) = similarity (A2*\<^sub>m\<^sub>mA1) M" using assms unfolding similarity_def by (simp add: mult_mm_assoc mat_inv_mult_mm) lemma similarity_inv: assumes "similarity A M1 = M2" and "mat_det A \ 0" shows "similarity (mat_inv A) M2 = M1" using assms unfolding similarity_def by (metis mat_det_mult mult_mm_assoc mult_mm_inv_l mult_mm_inv_r mult_zero_left) end diff --git a/thys/Complex_Geometry/Moebius.thy b/thys/Complex_Geometry/Moebius.thy --- a/thys/Complex_Geometry/Moebius.thy +++ b/thys/Complex_Geometry/Moebius.thy @@ -1,1535 +1,1535 @@ (* -------------------------------------------------------------------------- *) section \Möbius transformations\ (* -------------------------------------------------------------------------- *) text \Möbius transformations (also called homographic, linear fractional, or bilinear transformations) are the fundamental transformations of the extended complex plane. Here they are introduced algebraically. Each transformation is represented by a regular (non-singular, non-degenerate) $2\times 2$ matrix that acts linearly on homogeneous coordinates. As proportional homogeneous coordinates represent same points of $\mathbb{\overline{C}}$, proportional matrices will represent the same Möbius transformation.\ theory Moebius imports Homogeneous_Coordinates begin (* -------------------------------------------------------------------------- *) subsection \Definition of Möbius transformations\ (* -------------------------------------------------------------------------- *) typedef moebius_mat = "{M::complex_mat. mat_det M \ 0}" by (rule_tac x="eye" in exI, simp) setup_lifting type_definition_moebius_mat definition moebius_cmat_eq :: "complex_mat \ complex_mat \ bool" where [simp]: "moebius_cmat_eq A B \ (\ k::complex. k \ 0 \ B = k *\<^sub>s\<^sub>m A)" lift_definition moebius_mat_eq :: "moebius_mat \ moebius_mat \ bool" is moebius_cmat_eq done lemma moebius_mat_eq_refl [simp]: shows "moebius_mat_eq x x" by transfer simp quotient_type moebius = moebius_mat / moebius_mat_eq proof (rule equivpI) show "reflp moebius_mat_eq" unfolding reflp_def by transfer auto next show "symp moebius_mat_eq" unfolding symp_def by transfer (auto simp add: symp_def, rule_tac x="1/k" in exI, simp) next show "transp moebius_mat_eq" unfolding transp_def by transfer (auto simp add: transp_def, rule_tac x="ka*k" in exI, simp) qed definition mk_moebius_cmat :: "complex \ complex \ complex \ complex \ complex_mat" where [simp]: "mk_moebius_cmat a b c d = (let M = (a, b, c, d) in if mat_det M \ 0 then M else eye)" lift_definition mk_moebius_mat :: "complex \ complex \ complex \ complex \ moebius_mat" is mk_moebius_cmat by simp lift_definition mk_moebius :: "complex \ complex \ complex \ complex \ moebius" is mk_moebius_mat done lemma ex_mk_moebius: shows "\ a b c d. M = mk_moebius a b c d \ mat_det (a, b, c, d) \ 0" proof (transfer, transfer) fix M :: complex_mat assume "mat_det M \ 0" obtain a b c d where "M = (a, b, c, d)" by (cases M, auto) hence "moebius_cmat_eq M (mk_moebius_cmat a b c d) \ mat_det (a, b, c, d) \ 0" using \mat_det M \ 0\ by auto (rule_tac x=1 in exI, simp) thus "\a b c d. moebius_cmat_eq M (mk_moebius_cmat a b c d) \ mat_det (a, b, c, d) \ 0" by blast qed (* -------------------------------------------------------------------------- *) subsection \Action on points\ (* -------------------------------------------------------------------------- *) text \Möbius transformations are given as the action of Möbius group on the points of the extended complex plane (in homogeneous coordinates).\ definition moebius_pt_cmat_cvec :: "complex_mat \ complex_vec \ complex_vec" where [simp]: "moebius_pt_cmat_cvec M z = M *\<^sub>m\<^sub>v z" lift_definition moebius_pt_mmat_hcoords :: "moebius_mat \ complex_homo_coords \ complex_homo_coords" is moebius_pt_cmat_cvec by auto algebra+ lift_definition moebius_pt :: "moebius \ complex_homo \ complex_homo" is moebius_pt_mmat_hcoords proof transfer fix M M' x x' assume "moebius_cmat_eq M M'" "x \\<^sub>v x'" thus "moebius_pt_cmat_cvec M x \\<^sub>v moebius_pt_cmat_cvec M' x'" by (cases "M", cases "x", auto simp add: field_simps) (rule_tac x="k*ka" in exI, simp) qed lemma bij_moebius_pt [simp]: shows "bij (moebius_pt M)" unfolding bij_def inj_on_def surj_def proof safe fix x y assume "moebius_pt M x = moebius_pt M y" thus "x = y" proof (transfer, transfer) fix M x y assume "mat_det M \ 0" "moebius_pt_cmat_cvec M x \\<^sub>v moebius_pt_cmat_cvec M y" thus "x \\<^sub>v y" using mult_sv_mv[of _ M x] mult_mv_inv[of _ M] unfolding moebius_pt_cmat_cvec_def by (metis complex_cvec_eq_def) qed next fix y show "\x. y = moebius_pt M x" proof (transfer, transfer) fix y :: complex_vec and M :: complex_mat assume *: "y \ vec_zero" "mat_det M \ 0" let ?iM = "mat_inv M" let ?x = "?iM *\<^sub>m\<^sub>v y" have "?x \ vec_zero" using * by (metis mat_det_mult mat_eye_r mat_inv_r mult_cancel_right1 mult_mv_nonzero) moreover have "y \\<^sub>v moebius_pt_cmat_cvec M ?x" by (simp del: eye_def add: mat_inv_r[OF \mat_det M \ 0\]) ultimately show "\x\{v. v \ vec_zero}. y \\<^sub>v moebius_pt_cmat_cvec M x" by (rule_tac x="?x" in bexI, simp_all) qed qed lemma moebius_pt_eq_I: assumes "moebius_pt M z1 = moebius_pt M z2" shows "z1 = z2" using assms using bij_moebius_pt[of M] unfolding bij_def inj_on_def by blast lemma moebius_pt_neq_I [simp]: assumes "z1 \ z2" shows "moebius_pt M z1 \ moebius_pt M z2" using assms by (auto simp add: moebius_pt_eq_I) definition is_moebius :: "(complex_homo \ complex_homo) \ bool" where "is_moebius f \ (\ M. f = moebius_pt M)" text \In the classic literature Möbius transformations are often expressed in the form $\frac{az+b}{cz+d}$. The following lemma shows that when restricted to finite points, the action of Möbius transformations is bilinear.\ lemma moebius_pt_bilinear: assumes "mat_det (a, b, c, d) \ 0" shows "moebius_pt (mk_moebius a b c d) z = (if z \ \\<^sub>h then ((of_complex a) *\<^sub>h z +\<^sub>h (of_complex b)) :\<^sub>h ((of_complex c) *\<^sub>h z +\<^sub>h (of_complex d)) else (of_complex a) :\<^sub>h (of_complex c))" unfolding divide_def using assms proof (transfer, transfer) fix a b c d :: complex and z :: complex_vec obtain z1 z2 where zz: "z = (z1, z2)" by (cases z, auto) assume *: "mat_det (a, b, c, d) \ 0" "z \ vec_zero" let ?oc = "of_complex_cvec" show "moebius_pt_cmat_cvec (mk_moebius_cmat a b c d) z \\<^sub>v (if \ z \\<^sub>v \\<^sub>v then (?oc a *\<^sub>v z +\<^sub>v ?oc b) *\<^sub>v reciprocal_cvec (?oc c *\<^sub>v z +\<^sub>v ?oc d) else ?oc a *\<^sub>v reciprocal_cvec (?oc c))" proof (cases "z \\<^sub>v \\<^sub>v") case True thus ?thesis using zz * by auto next case False hence "z2 \ 0" using zz inf_cvec_z2_zero_iff \z \ vec_zero\ by auto thus ?thesis using zz * False using regular_homogenous_system[of a b c d z1 z2] by auto qed qed (* -------------------------------------------------------------------------- *) subsection \Möbius group\ (* -------------------------------------------------------------------------- *) text \Möbius elements form a group under composition. This group is called the \emph{projective general linear group} and denoted by $PGL(2, \mathbb{C})$ (the group $SGL(2, \mathbb{C})$ containing elements with the determinant $1$ can also be considered).\ text \Identity Möbius transformation is represented by the identity matrix.\ definition id_moebius_cmat :: "complex_mat" where [simp]: "id_moebius_cmat = eye" lift_definition id_moebius_mmat :: "moebius_mat" is id_moebius_cmat by simp lift_definition id_moebius :: "moebius" is id_moebius_mmat done lemma moebius_pt_moebius_id [simp]: shows "moebius_pt id_moebius = id" unfolding id_def apply (rule ext, transfer, transfer) using eye_mv_l by simp lemma mk_moeibus_id [simp]: shows "mk_moebius a 0 0 a = id_moebius" by (transfer, transfer, simp) text \The inverse Möbius transformation is obtained by taking the inverse representative matrix.\ definition moebius_inv_cmat :: "complex_mat \ complex_mat" where [simp]: "moebius_inv_cmat M = mat_inv M" lift_definition moebius_inv_mmat :: "moebius_mat \ moebius_mat" is moebius_inv_cmat by (simp add: mat_det_inv) lift_definition moebius_inv :: "moebius \ moebius" is "moebius_inv_mmat" proof (transfer) fix x y assume "moebius_cmat_eq x y" thus "moebius_cmat_eq (moebius_inv_cmat x) (moebius_inv_cmat y)" by (auto simp add: mat_inv_mult_sm) (rule_tac x="1/k" in exI, simp) qed lemma moebius_inv: shows "moebius_pt (moebius_inv M) = inv (moebius_pt M)" proof (rule inv_equality[symmetric]) fix x show "moebius_pt (moebius_inv M) (moebius_pt M x) = x" proof (transfer, transfer) fix M::complex_mat and x::complex_vec assume "mat_det M \ 0" "x \ vec_zero" show "moebius_pt_cmat_cvec (moebius_inv_cmat M) (moebius_pt_cmat_cvec M x) \\<^sub>v x" using eye_mv_l by (simp add: mat_inv_l[OF \mat_det M \ 0\]) qed next fix y show "moebius_pt M (moebius_pt (moebius_inv M) y) = y" proof (transfer, transfer) fix M::complex_mat and y::complex_vec assume "mat_det M \ 0" "y \ vec_zero" show "moebius_pt_cmat_cvec M (moebius_pt_cmat_cvec (moebius_inv_cmat M) y) \\<^sub>v y" using eye_mv_l by (simp add: mat_inv_r[OF \mat_det M \ 0\]) qed qed lemma is_moebius_inv [simp]: assumes "is_moebius m" shows "is_moebius (inv m)" using assms using moebius_inv unfolding is_moebius_def by metis lemma moebius_inv_mk_moebus [simp]: assumes "mat_det (a, b, c, d) \ 0" shows "moebius_inv (mk_moebius a b c d) = mk_moebius (d/(a*d-b*c)) (-b/(a*d-b*c)) (-c/(a*d-b*c)) (a/(a*d-b*c))" using assms by (transfer, transfer) (auto, rule_tac x=1 in exI, simp_all add: field_simps) text \Composition of Möbius elements is obtained by multiplying their representing matrices.\ definition moebius_comp_cmat :: "complex_mat \ complex_mat \ complex_mat" where [simp]: "moebius_comp_cmat M1 M2 = M1 *\<^sub>m\<^sub>m M2" lift_definition moebius_comp_mmat :: "moebius_mat \ moebius_mat \ moebius_mat" is moebius_comp_cmat by simp lift_definition moebius_comp :: "moebius \ moebius \ moebius" is moebius_comp_mmat by transfer (simp, (erule exE)+, rule_tac x="k*ka" in exI, simp add: field_simps) lemma moebius_comp: shows "moebius_pt (moebius_comp M1 M2) = moebius_pt M1 \ moebius_pt M2" unfolding comp_def by (rule ext, transfer, transfer, simp) lemma moebius_pt_comp [simp]: shows "moebius_pt (moebius_comp M1 M2) z = moebius_pt M1 (moebius_pt M2 z)" by (auto simp add: moebius_comp) lemma is_moebius_comp [simp]: assumes "is_moebius m1" and "is_moebius m2" shows "is_moebius (m1 \ m2)" using assms unfolding is_moebius_def using moebius_comp by metis lemma moebius_comp_mk_moebius [simp]: assumes "mat_det (a, b, c, d) \ 0" and "mat_det (a', b', c', d') \ 0" shows "moebius_comp (mk_moebius a b c d) (mk_moebius a' b' c' d') = mk_moebius (a * a' + b * c') (a * b' + b * d') (c * a' + d * c') (c * b' + d * d')" using mat_det_mult[of "(a, b, c, d)" "(a', b', c', d')"] using assms by (transfer, transfer) (auto, rule_tac x=1 in exI, simp) instantiation moebius :: group_add begin definition plus_moebius :: "moebius \ moebius \ moebius" where [simp]: "plus_moebius = moebius_comp" definition uminus_moebius :: "moebius \ moebius" where [simp]: "uminus_moebius = moebius_inv" definition zero_moebius :: "moebius" where [simp]: "zero_moebius = id_moebius" definition minus_moebius :: "moebius \ moebius \ moebius" where [simp]: "minus_moebius A B = A + (-B)" instance proof fix a b c :: moebius show "a + b + c = a + (b + c)" unfolding plus_moebius_def proof (transfer, transfer) fix a b c :: complex_mat assume "mat_det a \ 0" "mat_det b \ 0" "mat_det c \ 0" show "moebius_cmat_eq (moebius_comp_cmat (moebius_comp_cmat a b) c) (moebius_comp_cmat a (moebius_comp_cmat b c))" by simp (rule_tac x="1" in exI, simp add: mult_mm_assoc) qed next fix a :: moebius show "a + 0 = a" unfolding plus_moebius_def zero_moebius_def proof (transfer, transfer) fix A :: complex_mat assume "mat_det A \ 0" thus "moebius_cmat_eq (moebius_comp_cmat A id_moebius_cmat) A" using mat_eye_r by simp qed next fix a :: moebius show "0 + a = a" unfolding plus_moebius_def zero_moebius_def proof (transfer, transfer) fix A :: complex_mat assume "mat_det A \ 0" thus "moebius_cmat_eq (moebius_comp_cmat id_moebius_cmat A) A" using mat_eye_l by simp qed next fix a :: moebius show "- a + a = 0" unfolding plus_moebius_def uminus_moebius_def zero_moebius_def proof (transfer, transfer) fix a :: complex_mat assume "mat_det a \ 0" thus "moebius_cmat_eq (moebius_comp_cmat (moebius_inv_cmat a) a) id_moebius_cmat" by (simp add: mat_inv_l) qed next fix a b :: moebius show "a + - b = a - b" unfolding minus_moebius_def by simp qed end text \Composition with inverse\ lemma moebius_comp_inv_left [simp]: shows "moebius_comp (moebius_inv M) M = id_moebius" by (metis left_minus plus_moebius_def uminus_moebius_def zero_moebius_def) lemma moebius_comp_inv_right [simp]: shows "moebius_comp M (moebius_inv M) = id_moebius" by (metis right_minus plus_moebius_def uminus_moebius_def zero_moebius_def) lemma moebius_pt_comp_inv_left [simp]: shows "moebius_pt (moebius_inv M) (moebius_pt M z) = z" by (subst moebius_pt_comp[symmetric], simp) lemma moebius_pt_comp_inv_right [simp]: shows "moebius_pt M (moebius_pt (moebius_inv M) z) = z" by (subst moebius_pt_comp[symmetric], simp) lemma moebius_pt_comp_inv_image_left [simp]: shows "moebius_pt (moebius_inv M) ` moebius_pt M ` A = A" by force lemma moebius_pt_comp_inv_image_right [simp]: shows "moebius_pt M ` moebius_pt (moebius_inv M) ` A = A" by force lemma moebius_pt_invert: assumes "moebius_pt M z1 = z2" shows "moebius_pt (moebius_inv M) z2 = z1" using assms[symmetric] by simp lemma moebius_pt_moebius_inv_in_set [simp]: assumes "moebius_pt M z \ A" shows "z \ moebius_pt (moebius_inv M) ` A" using assms using image_iff by fastforce (* -------------------------------------------------------------------------- *) subsection \Special kinds of Möbius transformations\ (* -------------------------------------------------------------------------- *) (* -------------------------------------------------------------------------- *) subsubsection \Reciprocal (1/z) as a Möbius transformation\ (* -------------------------------------------------------------------------- *) definition moebius_reciprocal :: "moebius" where "moebius_reciprocal = mk_moebius 0 1 1 0" lemma moebius_reciprocal [simp]: shows "moebius_pt moebius_reciprocal = reciprocal" unfolding moebius_reciprocal_def by (rule ext, transfer, transfer) (force simp add: split_def) lemma moebius_reciprocal_inv [simp]: shows "moebius_inv moebius_reciprocal = moebius_reciprocal" unfolding moebius_reciprocal_def by (transfer, transfer) simp (* -------------------------------------------------------------------------- *) subsubsection \Euclidean similarities as a Möbius transform\ (* -------------------------------------------------------------------------- *) text\Euclidean similarities include Euclidean isometries (translations and rotations) and dilatations.\ definition moebius_similarity :: "complex \ complex \ moebius" where "moebius_similarity a b = mk_moebius a b 0 1" lemma moebius_pt_moebius_similarity [simp]: assumes "a \ 0" shows "moebius_pt (moebius_similarity a b) z = (of_complex a) *\<^sub>h z +\<^sub>h (of_complex b)" unfolding moebius_similarity_def using assms using mult_inf_right[of "of_complex a"] by (subst moebius_pt_bilinear, auto) text \Their action is a linear transformation of $\mathbb{C}.$\ lemma moebius_pt_moebius_similarity': assumes "a \ 0" shows "moebius_pt (moebius_similarity a b) = (\ z. (of_complex a) *\<^sub>h z +\<^sub>h (of_complex b))" using moebius_pt_moebius_similarity[OF assms, symmetric] by simp lemma is_moebius_similarity': assumes "a \ 0\<^sub>h" and "a \ \\<^sub>h" and "b \ \\<^sub>h" shows "(\ z. a *\<^sub>h z +\<^sub>h b) = moebius_pt (moebius_similarity (to_complex a) (to_complex b))" proof- obtain ka kb where *: "a = of_complex ka" "ka \ 0" "b = of_complex kb" using assms using inf_or_of_complex[of a] inf_or_of_complex[of b] by auto thus ?thesis unfolding is_moebius_def using moebius_pt_moebius_similarity'[of ka kb] by simp qed lemma is_moebius_similarity: assumes "a \ 0\<^sub>h" and "a \ \\<^sub>h" and "b \ \\<^sub>h" shows "is_moebius (\ z. a *\<^sub>h z +\<^sub>h b)" using is_moebius_similarity'[OF assms] unfolding is_moebius_def by auto text \Euclidean similarities form a group.\ lemma moebius_similarity_id [simp]: shows "moebius_similarity 1 0 = id_moebius" unfolding moebius_similarity_def by simp lemma moebius_similarity_inv [simp]: assumes "a \ 0" shows "moebius_inv (moebius_similarity a b) = moebius_similarity (1/a) (-b/a)" using assms unfolding moebius_similarity_def by simp lemma moebius_similarity_uminus [simp]: assumes "a \ 0" shows "- moebius_similarity a b = moebius_similarity (1/a) (-b/a)" using assms by simp lemma moebius_similarity_comp [simp]: assumes "a \ 0" and "c \ 0" shows "moebius_comp (moebius_similarity a b) (moebius_similarity c d) = moebius_similarity (a*c) (a*d+b)" using assms unfolding moebius_similarity_def by simp lemma moebius_similarity_plus [simp]: assumes "a \ 0" and "c \ 0" shows "moebius_similarity a b + moebius_similarity c d = moebius_similarity (a*c) (a*d+b)" using assms by simp text \Euclidean similarities are the only Möbius group elements such that their action leaves the $\infty_{h}$ fixed.\ lemma moebius_similarity_inf [simp]: assumes "a \ 0" shows "moebius_pt (moebius_similarity a b) \\<^sub>h = \\<^sub>h" using assms unfolding moebius_similarity_def by (transfer, transfer, simp) lemma moebius_similarity_only_inf_to_inf: assumes "a \ 0" "moebius_pt (moebius_similarity a b) z = \\<^sub>h" shows "z = \\<^sub>h" using assms using inf_or_of_complex[of z] by auto lemma moebius_similarity_inf_iff [simp]: assumes "a \ 0" shows "moebius_pt (moebius_similarity a b) z = \\<^sub>h \ z = \\<^sub>h" using assms using moebius_similarity_only_inf_to_inf[of a b z] by auto lemma inf_fixed_only_moebius_similarity: assumes "moebius_pt M \\<^sub>h = \\<^sub>h" shows "\ a b. a \ 0 \ M = moebius_similarity a b" using assms unfolding moebius_similarity_def proof (transfer, transfer) fix M :: complex_mat obtain a b c d where MM: "M = (a, b, c, d)" by (cases M, auto) assume "mat_det M \ 0" "moebius_pt_cmat_cvec M \\<^sub>v \\<^sub>v \\<^sub>v" hence *: "c = 0" "a \ 0 \ d \ 0" using MM by auto show "\a b. a \ 0 \ moebius_cmat_eq M (mk_moebius_cmat a b 0 1)" proof (rule_tac x="a/d" in exI, rule_tac x="b/d" in exI) show "a/d \ 0 \ moebius_cmat_eq M (mk_moebius_cmat (a / d) (b / d) 0 1)" using MM * by simp (rule_tac x="1/d" in exI, simp) qed qed text \Euclidean similarities include translations, rotations, and dilatations.\ (* -------------------------------------------------------------------------- *) subsubsection \Translation\ (* -------------------------------------------------------------------------- *) definition moebius_translation where "moebius_translation v = moebius_similarity 1 v" lemma moebius_translation_comp [simp]: shows "moebius_comp (moebius_translation v1) (moebius_translation v2) = moebius_translation (v1 + v2)" unfolding moebius_translation_def by (simp add: field_simps) lemma moebius_translation_plus [simp]: shows "(moebius_translation v1) + (moebius_translation v2) = moebius_translation (v1 + v2)" by simp lemma moebius_translation_zero [simp]: shows "moebius_translation 0 = id_moebius" unfolding moebius_translation_def moebius_similarity_id by simp lemma moebius_translation_inv [simp]: shows "moebius_inv (moebius_translation v1) = moebius_translation (-v1)" using moebius_translation_comp[of v1 "-v1"] moebius_translation_zero using minus_unique[of "moebius_translation v1" "moebius_translation (-v1)"] by simp lemma moebius_translation_uminus [simp]: shows "- (moebius_translation v1) = moebius_translation (-v1)" by simp lemma moebius_translation_inv_translation [simp]: shows "moebius_pt (moebius_translation v) (moebius_pt (moebius_translation (-v)) z) = z" using moebius_translation_inv[symmetric, of v] by (simp del: moebius_translation_inv) lemma moebius_inv_translation_translation [simp]: shows "moebius_pt (moebius_translation (-v)) (moebius_pt (moebius_translation v) z) = z" using moebius_translation_inv[symmetric, of v] by (simp del: moebius_translation_inv) lemma moebius_pt_moebius_translation [simp]: shows "moebius_pt (moebius_translation v) (of_complex z) = of_complex (z + v)" unfolding moebius_translation_def by (simp add: field_simps) lemma moebius_pt_moebius_translation_inf [simp]: shows "moebius_pt (moebius_translation v) \\<^sub>h = \\<^sub>h" unfolding moebius_translation_def by simp (* -------------------------------------------------------------------------- *) subsubsection \Rotation\ (* -------------------------------------------------------------------------- *) definition moebius_rotation where "moebius_rotation \ = moebius_similarity (cis \) 0" lemma moebius_rotation_comp [simp]: shows "moebius_comp (moebius_rotation \1) (moebius_rotation \2) = moebius_rotation (\1 + \2)" unfolding moebius_rotation_def using moebius_similarity_comp[of "cis \1" "cis \2" 0 0] by (simp add: cis_mult) lemma moebius_rotation_plus [simp]: shows "(moebius_rotation \1) + (moebius_rotation \2) = moebius_rotation (\1 + \2)" by simp lemma moebius_rotation_zero [simp]: shows "moebius_rotation 0 = id_moebius" unfolding moebius_rotation_def using moebius_similarity_id by simp lemma moebius_rotation_inv [simp]: shows "moebius_inv (moebius_rotation \) = moebius_rotation (- \)" using moebius_rotation_comp[of \ "-\"] moebius_rotation_zero using minus_unique[of "moebius_rotation \" "moebius_rotation (-\)"] by simp lemma moebius_rotation_uminus [simp]: shows "- (moebius_rotation \) = moebius_rotation (- \)" by simp lemma moebius_rotation_inv_rotation [simp]: shows "moebius_pt (moebius_rotation \) (moebius_pt (moebius_rotation (-\)) z) = z" using moebius_rotation_inv[symmetric, of \] by (simp del: moebius_rotation_inv) lemma moebius_inv_rotation_rotation [simp]: shows "moebius_pt (moebius_rotation (-\)) (moebius_pt (moebius_rotation \) z) = z" using moebius_rotation_inv[symmetric, of \] by (simp del: moebius_rotation_inv) lemma moebius_pt_moebius_rotation [simp]: shows "moebius_pt (moebius_rotation \) (of_complex z) = of_complex (cis \ * z)" unfolding moebius_rotation_def by simp lemma moebius_pt_moebius_rotation_inf [simp]: shows "moebius_pt (moebius_rotation v) \\<^sub>h = \\<^sub>h" unfolding moebius_rotation_def by simp lemma moebius_pt_rotation_inf_iff [simp]: shows "moebius_pt (moebius_rotation v) x = \\<^sub>h \ x = \\<^sub>h" unfolding moebius_rotation_def using cis_neq_zero moebius_similarity_only_inf_to_inf by (simp del: moebius_pt_moebius_similarity) lemma moebius_pt_moebius_rotation_zero [simp]: shows "moebius_pt (moebius_rotation \) 0\<^sub>h = 0\<^sub>h" unfolding moebius_rotation_def by simp lemma moebius_pt_moebius_rotation_zero_iff [simp]: shows "moebius_pt (moebius_rotation \) x = 0\<^sub>h \ x = 0\<^sub>h" using moebius_pt_invert[of "moebius_rotation \" x "0\<^sub>h"] by auto lemma moebius_rotation_preserve_cmod [simp]: assumes "u \ \\<^sub>h" shows "cmod (to_complex (moebius_pt (moebius_rotation \) u)) = cmod (to_complex u)" using assms using inf_or_of_complex[of u] - by auto + by (auto simp: norm_mult) (* -------------------------------------------------------------------------- *) subsubsection \Dilatation\ (* -------------------------------------------------------------------------- *) definition moebius_dilatation where "moebius_dilatation a = moebius_similarity (cor a) 0" lemma moebius_dilatation_comp [simp]: assumes "a1 > 0" and "a2 > 0" shows "moebius_comp (moebius_dilatation a1) (moebius_dilatation a2) = moebius_dilatation (a1 * a2)" using assms unfolding moebius_dilatation_def by simp lemma moebius_dilatation_plus [simp]: assumes "a1 > 0" and "a2 > 0" shows "(moebius_dilatation a1) + (moebius_dilatation a2) = moebius_dilatation (a1 * a2)" using assms by simp lemma moebius_dilatation_zero [simp]: shows "moebius_dilatation 1 = id_moebius" unfolding moebius_dilatation_def using moebius_similarity_id by simp lemma moebius_dilatation_inverse [simp]: assumes "a > 0" shows "moebius_inv (moebius_dilatation a) = moebius_dilatation (1/a)" using assms unfolding moebius_dilatation_def by simp lemma moebius_dilatation_uminus [simp]: assumes "a > 0" shows "- (moebius_dilatation a) = moebius_dilatation (1/a)" using assms by simp lemma moebius_pt_dilatation [simp]: assumes "a \ 0" shows "moebius_pt (moebius_dilatation a) (of_complex z) = of_complex (cor a * z)" using assms unfolding moebius_dilatation_def by simp (* -------------------------------------------------------------------------- *) subsubsection \Rotation-dilatation\ (* -------------------------------------------------------------------------- *) definition moebius_rotation_dilatation where "moebius_rotation_dilatation a = moebius_similarity a 0" lemma moebius_rotation_dilatation: assumes "a \ 0" shows "moebius_rotation_dilatation a = moebius_rotation (arg a) + moebius_dilatation (cmod a)" using assms unfolding moebius_rotation_dilatation_def moebius_rotation_def moebius_dilatation_def by simp (* -------------------------------------------------------------------------- *) subsubsection \Conjugate Möbius\ (* -------------------------------------------------------------------------- *) text \Conjugation is not a Möbius transformation, and conjugate Möbius transformations (obtained by conjugating each matrix element) do not represent conjugation function (although they are somewhat related).\ lift_definition conjugate_moebius_mmat :: "moebius_mat \ moebius_mat" is mat_cnj by auto lift_definition conjugate_moebius :: "moebius \ moebius" is conjugate_moebius_mmat by transfer (auto simp add: mat_cnj_def) lemma conjugate_moebius: shows "conjugate \ moebius_pt M = moebius_pt (conjugate_moebius M) \ conjugate" apply (rule ext, simp) apply (transfer, transfer) using vec_cnj_mult_mv by auto (* -------------------------------------------------------------------------- *) subsection \Decomposition of M\"obius transformations\ (* -------------------------------------------------------------------------- *) text \Every Euclidean similarity can be decomposed using translations, rotations, and dilatations.\ lemma similarity_decomposition: assumes "a \ 0" shows "moebius_similarity a b = (moebius_translation b) + (moebius_rotation (arg a)) + (moebius_dilatation (cmod a))" proof- have "moebius_similarity a b = (moebius_translation b) + (moebius_rotation_dilatation a)" using assms unfolding moebius_rotation_dilatation_def moebius_translation_def moebius_similarity_def by auto thus ?thesis using moebius_rotation_dilatation [OF assms] by (auto simp add: add.assoc simp del: plus_moebius_def) qed text \A very important fact is that every Möbius transformation can be composed of Euclidean similarities and a reciprocation.\ lemma moebius_decomposition: assumes "c \ 0" and "a*d - b*c \ 0" shows "mk_moebius a b c d = moebius_translation (a/c) + moebius_rotation_dilatation ((b*c - a*d)/(c*c)) + moebius_reciprocal + moebius_translation (d/c)" using assms unfolding moebius_rotation_dilatation_def moebius_translation_def moebius_similarity_def plus_moebius_def moebius_reciprocal_def by (simp add: field_simps) (transfer, transfer, auto simp add: field_simps, rule_tac x="1/c" in exI, simp) lemma moebius_decomposition_similarity: assumes "a \ 0" shows "mk_moebius a b 0 d = moebius_similarity (a/d) (b/d)" using assms unfolding moebius_similarity_def by (transfer, transfer, auto, rule_tac x="1/d" in exI, simp) text \Decomposition is used in many proofs. Namely, to show that every Möbius transformation has some property, it suffices to show that reciprocation and all Euclidean similarities have that property, and that the property is preserved under compositions.\ lemma wlog_moebius_decomposition: assumes trans: "\ v. P (moebius_translation v)" and rot: "\ \. P (moebius_rotation \)" and dil: "\ k. P (moebius_dilatation k)" and recip: "P (moebius_reciprocal)" and comp: "\ M1 M2. \P M1; P M2\ \ P (M1 + M2)" shows "P M" proof- obtain a b c d where "M = mk_moebius a b c d" "mat_det (a, b, c, d) \ 0" using ex_mk_moebius[of M] by auto show ?thesis proof (cases "c = 0") case False show ?thesis using moebius_decomposition[of c a d b] \mat_det (a, b, c, d) \ 0\ \c \ 0\ \M = mk_moebius a b c d\ using moebius_rotation_dilatation[of "(b*c - a*d) / (c*c)"] using trans[of "a/c"] rot[of "arg ((b*c - a*d) / (c*c))"] dil[of "cmod ((b*c - a*d) / (c*c))"] recip using comp by (simp add: trans) next case True hence "M = moebius_similarity (a/d) (b/d)" using \M = mk_moebius a b c d\ \mat_det (a, b, c, d) \ 0\ using moebius_decomposition_similarity by auto thus ?thesis using \c = 0\ \mat_det (a, b, c, d) \ 0\ using similarity_decomposition[of "a/d" "b/d"] using trans[of "b/d"] rot[of "arg (a/d)"] dil[of "cmod (a/d)"] comp by simp qed qed (* -------------------------------------------------------------------------- *) subsection \Cross ratio and Möbius existence\ (* -------------------------------------------------------------------------- *) text \For any fixed three points $z1$, $z2$ and $z3$, @{term "cross_ratio z z1 z2 z3"} can be seen as a function of a single variable $z$.\ lemma is_moebius_cross_ratio: assumes "z1 \ z2" and "z2 \ z3" and "z1 \ z3" shows "is_moebius (\ z. cross_ratio z z1 z2 z3)" proof- have "\ M. \ z. cross_ratio z z1 z2 z3 = moebius_pt M z" using assms proof (transfer, transfer) fix z1 z2 z3 assume vz: "z1 \ vec_zero" "z2 \ vec_zero" "z3 \ vec_zero" obtain z1' z1'' where zz1: "z1 = (z1', z1'')" by (cases z1, auto) obtain z2' z2'' where zz2: "z2 = (z2', z2'')" by (cases z2, auto) obtain z3' z3'' where zz3: "z3 = (z3', z3'')" by (cases z3, auto) let ?m23 = "z2'*z3''-z3'*z2''" let ?m21 = "z2'*z1''-z1'*z2''" let ?m13 = "z1'*z3''-z3'*z1''" let ?M = "(z1''*?m23, -z1'*?m23, z3''*?m21, -z3'*?m21)" assume "\ z1 \\<^sub>v z2" "\ z2 \\<^sub>v z3" "\ z1 \\<^sub>v z3" hence *: "?m23 \ 0" "?m21 \ 0" "?m13 \ 0" using vz zz1 zz2 zz3 using complex_cvec_eq_mix[of z1' z1'' z2' z2''] using complex_cvec_eq_mix[of z1' z1'' z3' z3''] using complex_cvec_eq_mix[of z2' z2'' z3' z3''] by (auto simp del: complex_cvec_eq_def simp add: field_simps) have "mat_det ?M = ?m21*?m23*?m13" by (simp add: field_simps) hence "mat_det ?M \ 0" using * by simp moreover have "\z\{v. v \ vec_zero}. cross_ratio_cvec z z1 z2 z3 \\<^sub>v moebius_pt_cmat_cvec ?M z" proof fix z assume "z \ {v. v \ vec_zero}" hence "z \ vec_zero" by simp obtain z' z'' where zz: "z = (z', z'')" by (cases z, auto) let ?m01 = "z'*z1''-z1'*z''" let ?m03 = "z'*z3''-z3'*z''" have "?m01 \ 0 \ ?m03 \ 0" proof (cases "z'' = 0 \ z1'' = 0 \ z3'' = 0") case True thus ?thesis using * \z \ vec_zero\ zz by auto next case False hence 1: "z'' \ 0 \ z1'' \ 0 \ z3'' \ 0" by simp show ?thesis proof (rule ccontr) assume "\ ?thesis" hence "z' * z1'' - z1' * z'' = 0" "z' * z3'' - z3' * z'' = 0" by auto hence "z1'/z1'' = z3'/z3''" using 1 zz \z \ vec_zero\ by (metis frac_eq_eq right_minus_eq) thus False using * 1 using frac_eq_eq by auto qed qed note * = * this show "cross_ratio_cvec z z1 z2 z3 \\<^sub>v moebius_pt_cmat_cvec ?M z" using * zz zz1 zz2 zz3 mult_mv_nonzero[of "z" ?M] \mat_det ?M \ 0\ by simp (rule_tac x="1" in exI, simp add: field_simps) qed ultimately show "\M\{M. mat_det M \ 0}. \z\{v. v \ vec_zero}. cross_ratio_cvec z z1 z2 z3 \\<^sub>v moebius_pt_cmat_cvec M z" by blast qed thus ?thesis by (auto simp add: is_moebius_def) qed text \Using properties of the cross-ratio, it is shown that there is a Möbius transformation mapping any three different points to $0_{hc}$, $1_{hc}$ and $\infty_{hc}$, respectively.\ lemma ex_moebius_01inf: assumes "z1 \ z2" and "z1 \ z3" and "z2 \ z3" shows "\ M. ((moebius_pt M z1 = 0\<^sub>h) \ (moebius_pt M z2 = 1\<^sub>h) \ (moebius_pt M z3 = \\<^sub>h))" using assms using is_moebius_cross_ratio[OF \z1 \ z2\ \z2 \ z3\ \z1 \ z3\] using cross_ratio_0[OF \z1 \ z2\ \z1 \ z3\] cross_ratio_1[OF \z1 \ z2\ \z2 \ z3\] cross_ratio_inf[OF \z1 \ z3\ \z2 \ z3\] by (metis is_moebius_def) text \There is a Möbius transformation mapping any three different points to any three different points.\ lemma ex_moebius: assumes "z1 \ z2" and "z1 \ z3" and "z2 \ z3" "w1 \ w2" and "w1 \ w3" and "w2 \ w3" shows "\ M. ((moebius_pt M z1 = w1) \ (moebius_pt M z2 = w2) \ (moebius_pt M z3 = w3))" proof- obtain M1 where *: "moebius_pt M1 z1 = 0\<^sub>h \ moebius_pt M1 z2 = 1\<^sub>h \ moebius_pt M1 z3 = \\<^sub>h" using ex_moebius_01inf[OF assms(1-3)] by auto obtain M2 where **: "moebius_pt M2 w1 = 0\<^sub>h \ moebius_pt M2 w2 = 1\<^sub>h \ moebius_pt M2 w3 = \\<^sub>h" using ex_moebius_01inf[OF assms(4-6)] by auto let ?M = "moebius_comp (moebius_inv M2) M1" show ?thesis using * ** by (rule_tac x="?M" in exI, auto simp add: moebius_pt_invert) qed lemma ex_moebius_1: shows "\ M. moebius_pt M z1 = w1" proof- obtain z2 z3 where "z1 \ z2" "z1 \ z3" "z2 \ z3" using ex_3_different_points[of z1] by auto moreover obtain w2 w3 where "w1 \ w2" "w1 \ w3" "w2 \ w3" using ex_3_different_points[of w1] by auto ultimately show ?thesis using ex_moebius[of z1 z2 z3 w1 w2 w3] by auto qed text \The next lemma turns out to have very important applications in further proof development, as it enables so called ,,without-loss-of-generality (wlog)'' reasoning \cite{wlog}. Namely, if the property is preserved under Möbius transformations, then instead of three arbitrary different points one can consider only the case of points $0_{hc}$, $1_{hc}$, and $\infty_{hc}$.\ lemma wlog_moebius_01inf: fixes M::moebius assumes "P 0\<^sub>h 1\<^sub>h \\<^sub>h" and "z1 \ z2" and "z2 \ z3" and "z1 \ z3" "\ M a b c. P a b c \ P (moebius_pt M a) (moebius_pt M b) (moebius_pt M c)" shows "P z1 z2 z3" proof- from assms obtain M where *: "moebius_pt M z1 = 0\<^sub>h" "moebius_pt M z2 = 1\<^sub>h" "moebius_pt M z3 = \\<^sub>h" using ex_moebius_01inf[of z1 z2 z3] by auto have **: "moebius_pt (moebius_inv M) 0\<^sub>h = z1" "moebius_pt (moebius_inv M) 1\<^sub>h = z2" "moebius_pt (moebius_inv M) \\<^sub>h = z3" by (subst *[symmetric], simp)+ thus ?thesis using assms by auto qed (* -------------------------------------------------------------------------- *) subsection \Fixed points and Möbius transformation uniqueness\ (* -------------------------------------------------------------------------- *) lemma three_fixed_points_01inf: assumes "moebius_pt M 0\<^sub>h = 0\<^sub>h" and "moebius_pt M 1\<^sub>h = 1\<^sub>h" and "moebius_pt M \\<^sub>h = \\<^sub>h" shows "M = id_moebius" using assms by (transfer, transfer, auto) lemma three_fixed_points: assumes "z1 \ z2" and "z1 \ z3" and "z2 \ z3" assumes "moebius_pt M z1 = z1" and "moebius_pt M z2 = z2" and "moebius_pt M z3 = z3" shows "M = id_moebius" proof- from assms obtain M' where *: "moebius_pt M' z1 = 0\<^sub>h" "moebius_pt M' z2 = 1\<^sub>h" "moebius_pt M' z3 = \\<^sub>h" using ex_moebius_01inf[of z1 z2 z3] by auto have **: "moebius_pt (moebius_inv M') 0\<^sub>h = z1" "moebius_pt (moebius_inv M') 1\<^sub>h = z2" "moebius_pt (moebius_inv M') \\<^sub>h = z3" by (subst *[symmetric], simp)+ have "M' + M + (-M') = 0" unfolding zero_moebius_def apply (rule three_fixed_points_01inf) using * ** assms by (simp add: moebius_comp[symmetric])+ thus ?thesis by (metis eq_neg_iff_add_eq_0 minus_add_cancel zero_moebius_def) qed lemma unique_moebius_three_points: assumes "z1 \ z2" and "z1 \ z3" and "z2 \ z3" assumes "moebius_pt M1 z1 = w1" and "moebius_pt M1 z2 = w2" and "moebius_pt M1 z3 = w3" "moebius_pt M2 z1 = w1" and "moebius_pt M2 z2 = w2" and "moebius_pt M2 z3 = w3" shows "M1 = M2" proof- let ?M = "moebius_comp (moebius_inv M2) M1" have "moebius_pt ?M z1 = z1" using \moebius_pt M1 z1 = w1\ \moebius_pt M2 z1 = w1\ by (auto simp add: moebius_pt_invert) moreover have "moebius_pt ?M z2 = z2" using \moebius_pt M1 z2 = w2\ \moebius_pt M2 z2 = w2\ by (auto simp add: moebius_pt_invert) moreover have "moebius_pt ?M z3 = z3" using \moebius_pt M1 z3 = w3\ \moebius_pt M2 z3 = w3\ by (auto simp add: moebius_pt_invert) ultimately have "?M = id_moebius" using assms three_fixed_points by auto thus ?thesis by (metis add_minus_cancel left_minus plus_moebius_def uminus_moebius_def zero_moebius_def) qed text \There is a unique Möbius transformation mapping three different points to other three different points.\ lemma ex_unique_moebius_three_points: assumes "z1 \ z2" and "z1 \ z3" and "z2 \ z3" "w1 \ w2" and "w1 \ w3" and "w2 \ w3" shows "\! M. ((moebius_pt M z1 = w1) \ (moebius_pt M z2 = w2) \ (moebius_pt M z3 = w3))" proof- obtain M where *: "moebius_pt M z1 = w1 \ moebius_pt M z2 = w2 \ moebius_pt M z3 = w3" using ex_moebius[OF assms] by auto show ?thesis unfolding Ex1_def proof (rule_tac x="M" in exI, rule) show "\y. moebius_pt y z1 = w1 \ moebius_pt y z2 = w2 \ moebius_pt y z3 = w3 \ y = M" using * using unique_moebius_three_points[OF assms(1-3)] by simp qed (simp add: *) qed lemma ex_unique_moebius_three_points_fun: assumes "z1 \ z2" and "z1 \ z3" and "z2 \ z3" "w1 \ w2" and "w1 \ w3" and "w2 \ w3" shows "\! f. is_moebius f \ (f z1 = w1) \ (f z2 = w2) \ (f z3 = w3)" proof- obtain M where "moebius_pt M z1 = w1" "moebius_pt M z2 = w2" "moebius_pt M z3 = w3" using ex_unique_moebius_three_points[OF assms] by auto thus ?thesis using ex_unique_moebius_three_points[OF assms] unfolding Ex1_def by (rule_tac x="moebius_pt M" in exI) (auto simp add: is_moebius_def) qed text \Different Möbius transformations produce different actions.\ lemma unique_moebius_pt: assumes "moebius_pt M1 = moebius_pt M2" shows "M1 = M2" using assms unique_moebius_three_points[of "0\<^sub>h" "1\<^sub>h" "\\<^sub>h"] by auto lemma is_cross_ratio_01inf: assumes "z1 \ z2" and "z1 \ z3" and "z2 \ z3" and "is_moebius f" assumes "f z1 = 0\<^sub>h" and "f z2 = 1\<^sub>h" and "f z3 = \\<^sub>h" shows "f = (\ z. cross_ratio z z1 z2 z3)" using assms using cross_ratio_0[OF \z1 \ z2\ \z1 \ z3\] cross_ratio_1[OF \z1 \ z2\ \z2 \ z3\] cross_ratio_inf[OF \z1 \ z3\ \z2 \ z3\] using is_moebius_cross_ratio[OF \z1 \ z2\ \z2 \ z3\ \z1 \ z3\] using ex_unique_moebius_three_points_fun[OF \z1 \ z2\ \z1 \ z3\ \z2 \ z3\, of "0\<^sub>h" "1\<^sub>h" "\\<^sub>h"] by auto text \Möbius transformations preserve cross-ratio.\ lemma moebius_preserve_cross_ratio [simp]: assumes "z1 \ z2" and "z1 \ z3" and "z2 \ z3" shows "cross_ratio (moebius_pt M z) (moebius_pt M z1) (moebius_pt M z2) (moebius_pt M z3) = cross_ratio z z1 z2 z3" proof- let ?f = "\ z. cross_ratio z z1 z2 z3" let ?M = "moebius_pt M" let ?iM = "inv ?M" have "(?f \ ?iM) (?M z1) = 0\<^sub>h" using bij_moebius_pt[of M] cross_ratio_0[OF \z1 \ z2\ \z1 \ z3\] by (simp add: bij_def) moreover have "(?f \ ?iM) (?M z2) = 1\<^sub>h" using bij_moebius_pt[of M] cross_ratio_1[OF \z1 \ z2\ \z2 \ z3\] by (simp add: bij_def) moreover have "(?f \ ?iM) (?M z3) = \\<^sub>h" using bij_moebius_pt[of M] cross_ratio_inf[OF \z1 \ z3\ \z2 \ z3\] by (simp add: bij_def) moreover have "is_moebius (?f \ ?iM)" by (rule is_moebius_comp, rule is_moebius_cross_ratio[OF \z1 \ z2\ \z2 \ z3\ \z1 \ z3\], rule is_moebius_inv, auto simp add: is_moebius_def) moreover have "?M z1 \ ?M z2" "?M z1 \ ?M z3" "?M z2 \ ?M z3" using assms by simp_all ultimately have "?f \ ?iM = (\ z. cross_ratio z (?M z1) (?M z2) (?M z3))" using assms using is_cross_ratio_01inf[of "?M z1" "?M z2" "?M z3" "?f \ ?iM"] by simp moreover have "(?f \ ?iM) (?M z) = cross_ratio z z1 z2 z3" using bij_moebius_pt[of M] by (simp add: bij_def) moreover have "(\ z. cross_ratio z (?M z1) (?M z2) (?M z3)) (?M z) = cross_ratio (?M z) (?M z1) (?M z2) (?M z3)" by simp ultimately show ?thesis by simp qed lemma conjugate_cross_ratio [simp]: assumes "z1 \ z2" and "z1 \ z3" and "z2 \ z3" shows "cross_ratio (conjugate z) (conjugate z1) (conjugate z2) (conjugate z3) = conjugate (cross_ratio z z1 z2 z3)" proof- let ?f = "\ z. cross_ratio z z1 z2 z3" let ?M = "conjugate" let ?iM = "conjugate" have "(conjugate \ ?f \ ?iM) (?M z1) = 0\<^sub>h" using cross_ratio_0[OF \z1 \ z2\ \z1 \ z3\] by simp moreover have "(conjugate \ ?f \ ?iM) (?M z2) = 1\<^sub>h" using cross_ratio_1[OF \z1 \ z2\ \z2 \ z3\] by simp moreover have "(conjugate \ ?f \ ?iM) (?M z3) = \\<^sub>h" using cross_ratio_inf[OF \z1 \ z3\ \z2 \ z3\] by simp moreover have "is_moebius (conjugate \ ?f \ ?iM)" proof- obtain M where "?f = moebius_pt M" using is_moebius_cross_ratio[OF \z1 \ z2\ \z2 \ z3\ \z1 \ z3\] by (auto simp add: is_moebius_def) thus ?thesis using conjugate_moebius[of M] by (auto simp add: comp_assoc is_moebius_def) qed moreover have "?M z1 \ ?M z2" "?M z1 \ ?M z3" "?M z2 \ ?M z3" using assms by (auto simp add: conjugate_inj) ultimately have "conjugate \ ?f \ ?iM = (\ z. cross_ratio z (?M z1) (?M z2) (?M z3))" using assms using is_cross_ratio_01inf[of "?M z1" "?M z2" "?M z3" "conjugate \ ?f \ ?iM"] by simp moreover have "(conjugate \ ?f \ ?iM) (?M z) = conjugate (cross_ratio z z1 z2 z3)" by simp moreover have "(\ z. cross_ratio z (?M z1) (?M z2) (?M z3)) (?M z) = cross_ratio (?M z) (?M z1) (?M z2) (?M z3)" by simp ultimately show ?thesis by simp qed lemma cross_ratio_reciprocal [simp]: assumes "u \ v" and "v \ w" and "u \ w" shows "cross_ratio (reciprocal z) (reciprocal u) (reciprocal v) (reciprocal w) = cross_ratio z u v w" using assms by (subst moebius_reciprocal[symmetric])+ (simp del: moebius_reciprocal) lemma cross_ratio_inversion [simp]: assumes "u \ v" and "v \ w" and "u \ w" shows "cross_ratio (inversion z) (inversion u) (inversion v) (inversion w) = conjugate (cross_ratio z u v w)" proof- have "reciprocal u \ reciprocal v" "reciprocal u \ reciprocal w" "reciprocal v \ reciprocal w" using assms by ((subst moebius_reciprocal[symmetric])+, simp del: moebius_reciprocal)+ thus ?thesis using assms unfolding inversion_def by simp qed lemma fixed_points_0inf': assumes "moebius_pt M 0\<^sub>h = 0\<^sub>h" and "moebius_pt M \\<^sub>h = \\<^sub>h" shows "\ k::complex_homo. (k \ 0\<^sub>h \ k \ \\<^sub>h) \ (\ z. moebius_pt M z = k *\<^sub>h z)" using assms proof (transfer, transfer) fix M :: complex_mat assume "mat_det M \ 0" obtain a b c d where MM: "M = (a, b, c, d)" by (cases M) auto assume "moebius_pt_cmat_cvec M 0\<^sub>v \\<^sub>v 0\<^sub>v" "moebius_pt_cmat_cvec M \\<^sub>v \\<^sub>v \\<^sub>v" hence *: "b = 0" "c = 0" "a \ 0 \ d \ 0" using MM by auto let ?z = "(a, d)" have "?z \ vec_zero" using * by simp moreover have "\ ?z \\<^sub>v 0\<^sub>v \ \ ?z \\<^sub>v \\<^sub>v" using * by simp moreover have "\z\{v. v \ vec_zero}. moebius_pt_cmat_cvec M z \\<^sub>v ?z *\<^sub>v z" using MM \mat_det M \ 0\ * by force ultimately show "\k\{v. v \ vec_zero}. (\ k \\<^sub>v 0\<^sub>v \ \ k \\<^sub>v \\<^sub>v) \ (\z\{v. v \ vec_zero}. moebius_pt_cmat_cvec M z \\<^sub>v k *\<^sub>v z)" by blast qed lemma fixed_points_0inf: assumes "moebius_pt M 0\<^sub>h = 0\<^sub>h" and "moebius_pt M \\<^sub>h = \\<^sub>h" shows "\ k::complex_homo. (k \ 0\<^sub>h \ k \ \\<^sub>h) \ moebius_pt M = (\ z. k *\<^sub>h z)" using fixed_points_0inf'[OF assms] by auto lemma ex_cross_ratio: assumes "u \ v" and "u \ w" and "v \ w" shows "\ z. cross_ratio z u v w = c" proof- obtain M where "(\ z. cross_ratio z u v w) = moebius_pt M" using assms is_moebius_cross_ratio[of u v w] unfolding is_moebius_def by auto hence *: "\ z. cross_ratio z u v w = moebius_pt M z" by metis let ?z = "moebius_pt (-M) c" have "cross_ratio ?z u v w = c" using * by auto thus ?thesis by auto qed lemma unique_cross_ratio: assumes "u \ v" and "v \ w" and "u \ w" assumes "cross_ratio z u v w = cross_ratio z' u v w" shows "z = z'" proof- obtain M where "(\ z. cross_ratio z u v w) = moebius_pt M" using is_moebius_cross_ratio[OF assms(1-3)] unfolding is_moebius_def by auto hence "moebius_pt M z = moebius_pt M z'" using assms(4) by metis thus ?thesis using moebius_pt_eq_I by metis qed lemma ex1_cross_ratio: assumes "u \ v" and "u \ w" and "v \ w" shows "\! z. cross_ratio z u v w = c" using assms ex_cross_ratio[OF assms, of c] unique_cross_ratio[of u v w] by blast (* -------------------------------------------------------------------------- *) subsection \Pole\ (* -------------------------------------------------------------------------- *) definition is_pole :: "moebius \ complex_homo \ bool" where "is_pole M z \ moebius_pt M z = \\<^sub>h" lemma ex1_pole: shows "\! z. is_pole M z" using bij_moebius_pt[of M] unfolding is_pole_def bij_def inj_on_def surj_def unfolding Ex1_def by (metis UNIV_I) definition pole :: "moebius \ complex_homo" where "pole M = (THE z. is_pole M z)" lemma pole_mk_moebius: assumes "is_pole (mk_moebius a b c d) z" and "c \ 0" and "a*d - b*c \ 0" shows "z = of_complex (-d/c)" proof- let ?t1 = "moebius_translation (a / c)" let ?rd = "moebius_rotation_dilatation ((b * c - a * d) / (c * c))" let ?r = "moebius_reciprocal" let ?t2 = "moebius_translation (d / c)" have "moebius_pt (?rd + ?r + ?t2) z = \\<^sub>h" using assms unfolding is_pole_def apply (subst (asm) moebius_decomposition) apply (auto simp add: moebius_comp[symmetric] moebius_translation_def) apply (subst moebius_similarity_only_inf_to_inf[of 1 "a/c"], auto) done hence "moebius_pt (?r + ?t2) z = \\<^sub>h" using \a*d - b*c \ 0\ \c \ 0\ unfolding moebius_rotation_dilatation_def by (simp del: moebius_pt_moebius_similarity) hence "moebius_pt ?t2 z = 0\<^sub>h" by simp thus ?thesis using moebius_pt_invert[of ?t2 z "0\<^sub>h"] by simp ((subst (asm) of_complex_zero[symmetric])+, simp del: of_complex_zero) qed lemma pole_similarity: assumes "is_pole (moebius_similarity a b) z" and "a \ 0" shows "z = \\<^sub>h" using assms unfolding is_pole_def using moebius_similarity_only_inf_to_inf[of a b z] by simp (* -------------------------------------------------------------------------- *) subsection \Homographies and antihomographies\ (* -------------------------------------------------------------------------- *) text \Inversion is not a Möbius transformation (it is a canonical example of so called anti-Möbius transformations, or antihomographies). All antihomographies are compositions of homographies and conjugation. The fundamental theorem of projective geometry (that we shall not prove) states that all automorphisms (bijective functions that preserve the cross-ratio) of $\mathbb{C}P^1$ are either homographies or antihomographies.\ definition is_homography :: "(complex_homo \ complex_homo) \ bool" where "is_homography f \ is_moebius f" definition is_antihomography :: "(complex_homo \ complex_homo) \ bool" where "is_antihomography f \ (\ f'. is_moebius f' \ f = f' \ conjugate)" text \Conjugation is not a Möbius transformation, but is antihomograhpy.\ lemma not_moebius_conjugate: shows "\ is_moebius conjugate" proof assume "is_moebius conjugate" then obtain M where *: "moebius_pt M = conjugate" unfolding is_moebius_def by metis hence "moebius_pt M 0\<^sub>h = 0\<^sub>h" "moebius_pt M 1\<^sub>h = 1\<^sub>h" "moebius_pt M \\<^sub>h = \\<^sub>h" by auto hence "M = id_moebius" using three_fixed_points_01inf by auto hence "conjugate = id" using * by simp moreover have "conjugate ii\<^sub>h \ ii\<^sub>h" using of_complex_inj[of "\" "-\"] by (subst of_complex_ii[symmetric])+ (auto simp del: of_complex_ii) ultimately show False by simp qed lemma conjugation_is_antihomography[simp]: shows "is_antihomography conjugate" unfolding is_antihomography_def by (rule_tac x="id" in exI, metis fun.map_id0 id_apply is_moebius_def moebius_pt_moebius_id) lemma inversion_is_antihomography [simp]: shows "is_antihomography inversion" using moebius_reciprocal unfolding inversion_sym is_antihomography_def is_moebius_def by metis text \Functions cannot simultaneously be homographies and antihomographies - the disjunction is exclusive.\ lemma homography_antihomography_exclusive: assumes "is_antihomography f" shows "\ is_homography f" proof assume "is_homography f" then obtain M where "f = moebius_pt M" unfolding is_homography_def is_moebius_def by auto then obtain M' where "moebius_pt M = moebius_pt M' \ conjugate" using assms unfolding is_antihomography_def is_moebius_def by auto hence "conjugate = moebius_pt (-M') \ moebius_pt M" by auto hence "conjugate = moebius_pt (-M' + M)" by (simp add: moebius_comp) thus False using not_moebius_conjugate unfolding is_moebius_def by metis qed (* -------------------------------------------------------------------------- *) subsection \Classification of Möbius transformations\ (* -------------------------------------------------------------------------- *) text \Möbius transformations can be classified to parabolic, elliptic and loxodromic. We do not develop this part of the theory in depth.\ lemma similarity_scale_1: assumes "k \ 0" shows "similarity (k *\<^sub>s\<^sub>m I) M = similarity I M" using assms unfolding similarity_def using mat_inv_mult_sm[of k I] by simp lemma similarity_scale_2: shows "similarity I (k *\<^sub>s\<^sub>m M) = k *\<^sub>s\<^sub>m (similarity I M)" unfolding similarity_def by auto lemma mat_trace_mult_sm [simp]: shows "mat_trace (k *\<^sub>s\<^sub>m M) = k * mat_trace M" by (cases M) (simp add: field_simps) definition moebius_mb_cmat :: "complex_mat \ complex_mat \ complex_mat" where [simp]: "moebius_mb_cmat I M = similarity I M" lift_definition moebius_mb_mmat :: "moebius_mat \ moebius_mat \ moebius_mat" is moebius_mb_cmat by (simp add: similarity_def mat_det_inv) lift_definition moebius_mb :: "moebius \ moebius \ moebius" is moebius_mb_mmat proof transfer fix M M' I I' assume "moebius_cmat_eq M M'" "moebius_cmat_eq I I'" thus "moebius_cmat_eq (moebius_mb_cmat I M) (moebius_mb_cmat I' M')" by (auto simp add: similarity_scale_1 similarity_scale_2) qed definition similarity_invar_cmat :: "complex_mat \ complex" where [simp]: "similarity_invar_cmat M = (mat_trace M)\<^sup>2 / mat_det M - 4" lift_definition similarity_invar_mmat :: "moebius_mat \ complex" is similarity_invar_cmat done lift_definition similarity_invar :: "moebius \ complex" is similarity_invar_mmat by transfer (auto simp add: power2_eq_square field_simps) lemma similarity_invar_moeibus_mb: shows "similarity_invar (moebius_mb I M) = similarity_invar M" by (transfer, transfer, simp) definition similar :: "moebius \ moebius \ bool" where "similar M1 M2 \ (\ I. moebius_mb I M1 = M2)" lemma similar_refl [simp]: shows "similar M M" unfolding similar_def by (rule_tac x="id_moebius" in exI) (transfer, transfer, simp) lemma similar_sym: assumes "similar M1 M2" shows "similar M2 M1" proof- from assms obtain I where "M2 = moebius_mb I M1" unfolding similar_def by auto hence "M1 = moebius_mb (moebius_inv I) M2" proof (transfer, transfer) fix M2 I M1 assume "moebius_cmat_eq M2 (moebius_mb_cmat I M1)" "mat_det I \ 0" then obtain k where "k \ 0" "similarity I M1 = k *\<^sub>s\<^sub>m M2" by auto thus "moebius_cmat_eq M1 (moebius_mb_cmat (moebius_inv_cmat I) M2)" using similarity_inv[of I M1 "k *\<^sub>s\<^sub>m M2", OF _ \mat_det I \ 0\] by (auto simp add: similarity_scale_2) (rule_tac x="1/k" in exI, simp) qed thus ?thesis unfolding similar_def by auto qed lemma similar_trans: assumes "similar M1 M2" and "similar M2 M3" shows "similar M1 M3" proof- obtain I1 I2 where "moebius_mb I1 M1 = M2" "moebius_mb I2 M2 = M3" using assms by (auto simp add: similar_def) thus ?thesis unfolding similar_def proof (rule_tac x="moebius_comp I1 I2" in exI, transfer, transfer) fix I1 I2 M1 M2 M3 assume "moebius_cmat_eq (moebius_mb_cmat I1 M1) M2" "moebius_cmat_eq (moebius_mb_cmat I2 M2) M3" "mat_det I1 \ 0" "mat_det I2 \ 0" thus "moebius_cmat_eq (moebius_mb_cmat (moebius_comp_cmat I1 I2) M1) M3" by (auto simp add: similarity_scale_2) (rule_tac x="ka*k" in exI, simp) qed qed end diff --git a/thys/Complex_Geometry/More_Complex.thy b/thys/Complex_Geometry/More_Complex.thy --- a/thys/Complex_Geometry/More_Complex.thy +++ b/thys/Complex_Geometry/More_Complex.thy @@ -1,1264 +1,1089 @@ (* -------------------------------------------------------------------------- *) subsection \Library Additions for Complex Numbers\ (* -------------------------------------------------------------------------- *) text \Some additional lemmas about complex numbers.\ theory More_Complex imports Complex_Main More_Transcendental Canonical_Angle begin text \Conjugation and @{term cis}\ - + declare cis_cnj[simp] lemma rcis_cnj: shows "cnj a = rcis (cmod a) (- arg a)" by (subst rcis_cmod_arg[of a, symmetric]) (simp add: rcis_def) lemmas complex_cnj = complex_cnj_diff complex_cnj_mult complex_cnj_add complex_cnj_divide complex_cnj_minus text \Some properties for @{term complex_of_real}. Also, since it is often used in our formalization we abbreviate it to @{term cor}.\ abbreviation cor :: "real \ complex" where "cor \ complex_of_real" -lemma cor_neg_one [simp]: - shows "cor (-1) = -1" - by simp - -lemma neg_cor_neg_one [simp]: - shows "- cor (-1) = 1" - by simp - lemma cmod_cis [simp]: assumes "a \ 0" shows "cor (cmod a) * cis (arg a) = a" using assms by (metis rcis_cmod_arg rcis_def) lemma cis_cmod [simp]: assumes "a \ 0" shows "cis (arg a) * cor (cmod a) = a" using assms cmod_cis[of a] by (simp add: field_simps) -lemma cor_add: - shows "cor (a + b) = cor a + cor b" - by (rule of_real_add) - -lemma cor_mult: - shows "cor (a * b) = cor a * cor b" - by (rule of_real_mult) - lemma cor_squared: shows "(cor x)\<^sup>2 = cor (x\<^sup>2)" by (simp add: power2_eq_square) lemma cor_sqrt_mult_cor_sqrt [simp]: shows "cor (sqrt A) * cor (sqrt A) = cor \A\" - by (metis cor_mult real_sqrt_abs2 real_sqrt_mult) + by (metis of_real_mult real_sqrt_mult_self) lemma cor_eq_0: "cor x + \ * cor y = 0 \ x = 0 \ y = 0" by (metis Complex_eq Im_complex_of_real Im_i_times Re_complex_of_real add_cancel_left_left of_real_eq_0_iff plus_complex.sel(2) zero_complex.code) lemma one_plus_square_neq_zero [simp]: shows "1 + (cor x)\<^sup>2 \ 0" by (metis (hide_lams, no_types) of_real_1 of_real_add of_real_eq_0_iff of_real_power power_one sum_power2_eq_zero_iff zero_neq_one) text \Additional lemmas about @{term Complex} constructor. Following newer versions of Isabelle, these should be deprecated.\ lemma complex_real_two [simp]: shows "Complex 2 0 = 2" by (simp add: Complex_eq) lemma complex_double [simp]: shows "(Complex a b) * 2 = Complex (2*a) (2*b)" by (simp add: Complex_eq) lemma complex_half [simp]: shows "(Complex a b) / 2 = Complex (a/2) (b/2)" by (subst complex_eq_iff) auto lemma Complex_scale1: shows "Complex (a * b) (a * c) = cor a * Complex b c" unfolding complex_of_real_def unfolding Complex_eq by (auto simp add: field_simps) lemma Complex_scale2: shows "Complex (a * c) (b * c) = Complex a b * cor c" unfolding complex_of_real_def unfolding Complex_eq by (auto simp add: field_simps) lemma Complex_scale3: shows "Complex (a / b) (a / c) = cor a * Complex (1 / b) (1 / c)" unfolding complex_of_real_def unfolding Complex_eq by (auto simp add: field_simps) lemma Complex_scale4: shows "c \ 0 \ Complex (a / c) (b / c) = Complex a b / cor c" unfolding complex_of_real_def unfolding Complex_eq by (auto simp add: field_simps power2_eq_square) lemma Complex_Re_express_cnj: shows "Complex (Re z) 0 = (z + cnj z) / 2" by (cases z) (simp add: Complex_eq) lemma Complex_Im_express_cnj: shows "Complex 0 (Im z) = (z - cnj z)/2" by (cases z) (simp add: Complex_eq) text \Additional properties of @{term cmod}.\ lemma complex_mult_cnj_cmod: shows "z * cnj z = cor ((cmod z)\<^sup>2)" using complex_norm_square by auto lemma cmod_square: shows "(cmod z)\<^sup>2 = Re (z * cnj z)" using complex_mult_cnj_cmod[of z] by (simp add: power2_eq_square) lemma cor_cmod_power_4 [simp]: shows "cor (cmod z) ^ 4 = (z * cnj z)\<^sup>2" - by (metis complex_norm_square cor_squared numeral_times_numeral power2_eq_square semiring_norm(11) semiring_norm(13) semiring_normalization_rules(36)) + by (simp add: complex_mult_cnj_cmod) lemma cnjE: assumes "x \ 0" shows "cnj x = cor ((cmod x)\<^sup>2) / x" using complex_mult_cnj_cmod[of x] assms by (auto simp add: field_simps) -lemma cmod_mult [simp]: - shows "cmod (a * b) = cmod a * cmod b" - by (rule norm_mult) - -lemma cmod_divide [simp]: - shows "cmod (a / b) = cmod a / cmod b" - by (rule norm_divide) - lemma cmod_cor_divide [simp]: shows "cmod (z / cor k) = cmod z / \k\" - by auto + by (simp add: norm_divide) lemma cmod_mult_minus_left_distrib [simp]: shows "cmod (z*z1 - z*z2) = cmod z * cmod(z1 - z2)" - by (metis bounded_bilinear.diff_right bounded_bilinear_mult cmod_mult) + by (metis norm_mult right_diff_distrib) lemma cmod_eqI: assumes "z1 * cnj z1 = z2 * cnj z2" shows "cmod z1 = cmod z2" using assms by (subst complex_mod_sqrt_Re_mult_cnj)+ auto lemma cmod_eqE: assumes "cmod z1 = cmod z2" shows "z1 * cnj z1 = z2 * cnj z2" -proof- - from assms have "cor ((cmod z1)\<^sup>2) = cor ((cmod z2)\<^sup>2)" - by auto - thus ?thesis - using complex_mult_cnj_cmod - by auto -qed + by (simp add: assms complex_mult_cnj_cmod) lemma cmod_eq_one [simp]: shows "cmod a = 1 \ a*cnj a = 1" by (metis cmod_eqE cmod_eqI complex_cnj_one monoid_mult_class.mult.left_neutral norm_one) text \We introduce @{term is_real} (the imaginary part of complex number is zero) and @{term is_imag} (real part of complex number is zero) operators and prove some of their properties.\ abbreviation is_real where "is_real z \ Im z = 0" abbreviation is_imag where "is_imag z \ Re z = 0" lemma real_imag_0: assumes "is_real a" "is_imag a" shows "a = 0" using assms by (simp add: complex.expand) lemma complex_eq_if_Re_eq: assumes "is_real z1" and "is_real z2" shows "z1 = z2 \ Re z1 = Re z2" using assms by (cases z1, cases z2) auto lemma mult_reals [simp]: assumes "is_real a" and "is_real b" shows "is_real (a * b)" using assms by auto lemma div_reals [simp]: assumes "is_real a" and "is_real b" shows "is_real (a / b)" using assms by (simp add: complex_is_Real_iff) lemma complex_of_real_Re [simp]: assumes "is_real k" shows "cor (Re k) = k" using assms by (cases k) (auto simp add: complex_of_real_def) lemma cor_cmod_real: assumes "is_real a" shows "cor (cmod a) = a \ cor (cmod a) = -a" using assms unfolding cmod_def by (cases "Re a > 0") auto lemma eq_cnj_iff_real: shows "cnj z = z \ is_real z" by (cases z) (simp add: Complex_eq) lemma eq_minus_cnj_iff_imag: shows "cnj z = -z \ is_imag z" by (cases z) (simp add: Complex_eq) lemma Re_divide_real: assumes "is_real b" and "b \ 0" shows "Re (a / b) = (Re a) / (Re b)" using assms by (simp add: complex_is_Real_iff) lemma Re_mult_real: assumes "is_real a" shows "Re (a * b) = (Re a) * (Re b)" using assms by simp lemma Im_mult_real: assumes "is_real a" shows "Im (a * b) = (Re a) * (Im b)" using assms by simp lemma Im_divide_real: assumes "is_real b" and "b \ 0" shows "Im (a / b) = (Im a) / (Re b)" using assms by (simp add: complex_is_Real_iff) -lemma Re_half [simp]: - shows "Re (x / 2) = Re x / 2" - by (rule Re_divide_numeral) - -lemma Re_double [simp]: - shows "Re (2 * x) = 2 * Re x" - using Re_mult_real[of "2" x] - by simp - -lemma Im_half [simp]: - shows "Im (z / 2) = Im z / 2" - by (subst Im_divide_real, auto) - -lemma Im_double [simp]: - shows "Im (2 * z) = 2 * Im z" - using Im_mult_real[of "2" z] - by simp - lemma Re_sgn: assumes "is_real R" shows "Re (sgn R) = sgn (Re R)" using assms by (metis Re_sgn complex_of_real_Re norm_of_real real_sgn_eq) lemma is_real_div: assumes "b \ 0" shows "is_real (a / b) \ a*cnj b = b*cnj a" using assms by (metis complex_cnj_divide complex_cnj_zero_iff eq_cnj_iff_real frac_eq_eq mult.commute) lemma is_real_mult_real: assumes "is_real a" and "a \ 0" shows "is_real b \ is_real (a * b)" using assms by (cases a, auto simp add: Complex_eq) lemma Im_express_cnj: shows "Im z = (z - cnj z) / (2 * \)" by (simp add: complex_diff_cnj field_simps) lemma Re_express_cnj: shows "Re z = (z + cnj z) / 2" by (simp add: complex_add_cnj) text \Rotation of complex number for 90 degrees in the positive direction.\ abbreviation rot90 where "rot90 z \ Complex (-Im z) (Re z)" lemma rot90_ii: shows "rot90 z = z * \" by (metis Complex_mult_i complex_surj) text \With @{term cnj_mix} we introduce scalar product between complex vectors. This operation shows to be useful to succinctly express some conditions.\ abbreviation cnj_mix where "cnj_mix z1 z2 \ cnj z1 * z2 + z1 * cnj z2" abbreviation scalprod where "scalprod z1 z2 \ cnj_mix z1 z2 / 2" lemma cnj_mix_minus: shows "cnj z1*z2 - z1*cnj z2 = \ * cnj_mix (rot90 z1) z2" by (cases z1, cases z2) (simp add: Complex_eq field_simps) lemma cnj_mix_minus': shows "cnj z1*z2 - z1*cnj z2 = rot90 (cnj_mix (rot90 z1) z2)" by (cases z1, cases z2) (simp add: Complex_eq field_simps) lemma cnj_mix_real [simp]: shows "is_real (cnj_mix z1 z2)" by (cases z1, cases z2) simp lemma scalprod_real [simp]: shows "is_real (scalprod z1 z2)" using cnj_mix_real by simp text \Additional properties of @{term cis} function.\ lemma cis_minus_pi2 [simp]: shows "cis (-pi/2) = -\" by (simp add: cis_inverse[symmetric]) lemma cis_pi2_minus_x [simp]: shows "cis (pi/2 - x) = \ * cis(-x)" using cis_divide[of "pi/2" x, symmetric] using cis_divide[of 0 x, symmetric] by simp lemma cis_pm_pi [simp]: shows "cis (x - pi) = - cis x" and "cis (x + pi) = - cis x" by (simp add: cis.ctr complex_minus)+ lemma cis_times_cis_opposite [simp]: shows "cis \ * cis (- \) = 1" by (simp add: cis_mult) text \@{term cis} repeats only after $2k\pi$\ lemma cis_eq: assumes "cis a = cis b" shows "\ k::int. a - b = 2 * k * pi" using assms sin_cos_eq[of a b] using cis.sel[of a] cis.sel[of b] by (cases "cis a", cases "cis b") auto text \@{term cis} is injective on $(-\pi, \pi]$.\ lemma cis_inj: assumes "-pi < \" and "\ \ pi" and "-pi < \'" and "\' \ pi" assumes "cis \ = cis \'" shows "\ = \'" using assms by (metis arg_unique sgn_cis) text \@{term cis} of an angle combined with @{term cis} of the opposite angle\ lemma cis_diff_cis_opposite [simp]: shows "cis \ - cis (- \) = 2 * \ * sin \" using Im_express_cnj[of "cis \"] by simp lemma cis_opposite_diff_cis [simp]: shows "cis (-\) - cis (\) = - 2 * \ * sin \" using cis_diff_cis_opposite[of "-\"] by simp lemma cis_add_cis_opposite [simp]: shows "cis \ + cis (-\) = 2 * cos \" -proof- - have "2 * cos \ = (cis \ + cnj (cis \))" - using Re_express_cnj[of "cis \"] - by (simp add: field_simps) - thus ?thesis - by simp -qed + by (metis cis.sel(1) cis_cnj complex_add_cnj) text \@{term cis} equal to 1 or -1\ lemma cis_one [simp]: assumes "sin \ = 0" and "cos \ = 1" shows "cis \ = 1" using assms by (auto simp add: cis.ctr one_complex.code) lemma cis_minus_one [simp]: assumes "sin \ = 0" and "cos \ = -1" shows "cis \ = -1" using assms by (auto simp add: cis.ctr Complex_eq_neg_1) (* -------------------------------------------------------------------------- *) subsubsection \Additional properties of complex number argument\ (* -------------------------------------------------------------------------- *) text \@{term arg} of real numbers\ lemma is_real_arg1: assumes "arg z = 0 \ arg z = pi" shows "is_real z" using assms using rcis_cmod_arg[of z] Im_rcis[of "cmod z" "arg z"] by auto lemma is_real_arg2: assumes "is_real z" shows "arg z = 0 \ arg z = pi" proof (cases "z = 0") - case True - thus ?thesis - by (auto simp add: arg_zero) -next case False - hence "sin (arg z) = 0" - using assms rcis_cmod_arg[of z] Im_rcis[of "cmod z" "arg z"] - by auto thus ?thesis using arg_bounded[of z] - using sin_0_iff_canon - by simp -qed + by (smt (verit, best) Im_sgn assms cis.simps(2) cis_arg div_0 sin_zero_pi_iff) +qed (auto simp add: arg_zero) lemma arg_complex_of_real_positive [simp]: assumes "k > 0" shows "arg (cor k) = 0" proof- have "cos (arg (Complex k 0)) > 0" using assms using rcis_cmod_arg[of "Complex k 0"] Re_rcis[of "cmod (Complex k 0)" "arg (Complex k 0)"] - by (smt complex.sel(1) mult_nonneg_nonpos norm_ge_zero) + using cmod_eq_Re by force thus ?thesis using assms is_real_arg2[of "cor k"] unfolding complex_of_real_def by auto qed lemma arg_complex_of_real_negative [simp]: assumes "k < 0" shows "arg (cor k) = pi" proof- have "cos (arg (Complex k 0)) < 0" using \k < 0\ rcis_cmod_arg[of "Complex k 0"] Re_rcis[of "cmod (Complex k 0)" "arg (Complex k 0)"] - by (smt complex.sel(1) mult_nonneg_nonneg norm_ge_zero) + by (metis complex.sel(1) mult_less_0_iff norm_not_less_zero) thus ?thesis using assms is_real_arg2[of "cor k"] unfolding complex_of_real_def by auto qed lemma arg_0_iff: shows "z \ 0 \ arg z = 0 \ is_real z \ Re z > 0" by (smt arg_complex_of_real_negative arg_complex_of_real_positive arg_zero complex_of_real_Re is_real_arg1 pi_gt_zero zero_complex.simps) lemma arg_pi_iff: shows "arg z = pi \ is_real z \ Re z < 0" by (smt arg_complex_of_real_negative arg_complex_of_real_positive arg_zero complex_of_real_Re is_real_arg1 pi_gt_zero zero_complex.simps) text \@{term arg} of imaginary numbers\ lemma is_imag_arg1: assumes "arg z = pi/2 \ arg z = -pi/2" shows "is_imag z" using assms using rcis_cmod_arg[of z] Re_rcis[of "cmod z" "arg z"] by (metis cos_minus cos_pi_half minus_divide_left mult_eq_0_iff) lemma is_imag_arg2: assumes "is_imag z" and "z \ 0" shows "arg z = pi/2 \ arg z = -pi/2" -proof- - have "cos (arg z) = 0" - using assms - by (metis Re_rcis no_zero_divisors norm_eq_zero rcis_cmod_arg) - thus ?thesis - using arg_bounded[of z] - using cos_0_iff_canon[of "arg z"] - by simp -qed + using arg_bounded assms cos_0_iff_canon cos_arg_i_mult_zero by presburger lemma arg_complex_of_real_times_i_positive [simp]: assumes "k > 0" shows "arg (cor k * \) = pi / 2" proof- have "sin (arg (Complex 0 k)) > 0" using \k > 0\ rcis_cmod_arg[of "Complex 0 k"] Im_rcis[of "cmod (Complex 0 k)" "arg (Complex 0 k)"] by (smt complex.sel(2) mult_nonneg_nonpos norm_ge_zero) thus ?thesis using assms is_imag_arg2[of "cor k * \"] using arg_zero complex_of_real_i by force qed lemma arg_complex_of_real_times_i_negative [simp]: assumes "k < 0" shows "arg (cor k * \) = - pi / 2" proof- have "sin (arg (Complex 0 k)) < 0" using \k < 0\ rcis_cmod_arg[of "Complex 0 k"] Im_rcis[of "cmod (Complex 0 k)" "arg (Complex 0 k)"] - by (smt complex.sel(2) mult_nonneg_nonneg norm_ge_zero) + by (metis complex.sel(2) mult_less_0_iff norm_not_less_zero) thus ?thesis using assms is_imag_arg2[of "cor k * \"] using arg_zero complex_of_real_i[of k] by (smt complex.sel(1) sin_pi_half sin_zero) qed lemma arg_pi2_iff: shows "z \ 0 \ arg z = pi / 2 \ is_imag z \ Im z > 0" by (smt Im_rcis Re_i_times Re_rcis arcsin_minus_1 cos_pi_half divide_minus_left mult.commute mult_cancel_right1 rcis_cmod_arg is_imag_arg2 sin_arcsin sin_pi_half zero_less_mult_pos zero_less_norm_iff) lemma arg_minus_pi2_iff: shows "z \ 0 \ arg z = - pi / 2 \ is_imag z \ Im z < 0" by (smt arg_pi2_iff complex.expand divide_cancel_right pi_neq_zero is_imag_arg1 is_imag_arg2 zero_complex.simps(1) zero_complex.simps(2)) lemma arg_ii [simp]: shows "arg \ = pi/2" -proof- - have "\ = cis (arg \)" - using rcis_cmod_arg[of \] - by (simp add: rcis_def) - hence "cos (arg \) = 0" "sin (arg \) = 1" - by (metis cis.simps(1) imaginary_unit.simps(1), metis cis.simps(2) imaginary_unit.simps(2)) - thus ?thesis - using cos_0_iff_canon[of "arg \"] arg_bounded[of \] - by auto -qed + by (metis arg_pi2_iff imaginary_unit.sel zero_less_one) lemma arg_minus_ii [simp]: shows "arg (-\) = -pi/2" proof- have "-\ = cis (arg (- \))" using rcis_cmod_arg[of "-\"] by (simp add: rcis_def) hence "cos (arg (-\)) = 0" "sin (arg (-\)) = -1" using cis.simps[of "arg (-\)"] by auto thus ?thesis using cos_0_iff_canon[of "arg (-\)"] arg_bounded[of "-\"] by fastforce qed text \Argument is a canonical angle\ lemma canon_ang_arg: shows "\arg z\ = arg z" using canon_ang_id[of "arg z"] arg_bounded by simp lemma arg_cis: shows "arg (cis \) = \\\" -proof (rule canon_ang_eqI[symmetric]) - show "- pi < arg (cis \) \ arg (cis \) \ pi" - using arg_bounded - by simp -next - show "\ k::int. arg (cis \) - \ = 2*k*pi" - proof- - have "cis (arg (cis \)) = cis \" - using cis_arg[of "cis \"] - by auto - thus ?thesis - using cis_eq - by auto - qed -qed + using arg_unique canon_ang canon_ang_cos canon_ang_sin cis.ctr sgn_cis by presburger text \Cosine and sine of @{term arg}\ lemma cos_arg: assumes "z \ 0" shows "cos (arg z) = Re z / cmod z" by (metis Complex.Re_sgn cis.simps(1) assms cis_arg) lemma sin_arg: assumes "z \ 0" shows "sin (arg z) = Im z / cmod z" by (metis Complex.Im_sgn cis.simps(2) assms cis_arg) text \Argument of product\ lemma cis_arg_mult: assumes "z1 * z2 \ 0" shows "cis (arg (z1 * z2)) = cis (arg z1 + arg z2)" -proof- - have "z1 * z2 = cor (cmod z1) * cor (cmod z2) * cis (arg z1) * cis (arg z2)" - using rcis_cmod_arg[of z1, symmetric] rcis_cmod_arg[of z2, symmetric] - unfolding rcis_def - by algebra - hence "z1 * z2 = cor (cmod (z1 * z2)) * cis (arg z1 + arg z2)" - using cis_mult[of "arg z1" "arg z2"] - by auto - hence "cor (cmod (z1 * z2)) * cis (arg z1 + arg z2) = cor (cmod (z1 * z2)) * cis (arg (z1 * z2))" - using assms - using rcis_cmod_arg[of "z1*z2"] - unfolding rcis_def - by auto - thus ?thesis - using mult_cancel_left[of "cor (cmod (z1 * z2))" "cis (arg z1 + arg z2)" "cis (arg (z1 * z2))"] - using assms - by auto -qed + by (metis assms cis_arg cis_mult mult_eq_0_iff sgn_mult) lemma arg_mult_2kpi: assumes "z1 * z2 \ 0" shows "\ k::int. arg (z1 * z2) = arg z1 + arg z2 + 2*k*pi" proof- have "cis (arg (z1*z2)) = cis (arg z1 + arg z2)" by (rule cis_arg_mult[OF assms]) thus ?thesis using cis_eq[of "arg (z1*z2)" "arg z1 + arg z2"] by (auto simp add: field_simps) qed lemma arg_mult: assumes "z1 * z2 \ 0" shows "arg(z1 * z2) = \arg z1 + arg z2\" proof- obtain k::int where "arg(z1 * z2) = arg z1 + arg z2 + 2*k*pi" using arg_mult_2kpi[of z1 z2] using assms by auto hence "\arg(z1 * z2)\ = \arg z1 + arg z2\" using canon_ang_eq by(simp add:field_simps) thus ?thesis using canon_ang_arg[of "z1*z2"] by auto qed lemma arg_mult_real_positive [simp]: assumes "k > 0" shows "arg (cor k * z) = arg z" proof (cases "z = 0") - case True - thus ?thesis - by (auto simp add: arg_zero) -next case False thus ?thesis - using assms - using arg_mult[of "cor k" z] - by (auto simp add: canon_ang_arg) -qed + using arg_mult assms canon_ang_arg by force +qed (auto simp: arg_zero) lemma arg_mult_real_negative [simp]: assumes "k < 0" shows "arg (cor k * z) = arg (-z)" proof (cases "z = 0") - case True - thus ?thesis - by (auto simp add: arg_zero) -next case False thus ?thesis using assms - using arg_mult[of "cor k" z] - using arg_mult[of "-1" z] - using arg_complex_of_real_negative[of k] arg_complex_of_real_negative[of "-1"] - by auto -qed + by (metis arg_mult_real_positive minus_mult_commute neg_0_less_iff_less of_real_minus minus_minus) +qed (auto simp: arg_zero) lemma arg_div_real_positive [simp]: assumes "k > 0" shows "arg (z / cor k) = arg z" proof(cases "z = 0") case True thus ?thesis by auto next case False thus ?thesis using assms using arg_mult_real_positive[of "1/k" z] by auto qed lemma arg_div_real_negative [simp]: assumes "k < 0" shows "arg (z / cor k) = arg (-z)" proof(cases "z = 0") case True thus ?thesis by auto next case False thus ?thesis using assms using arg_mult_real_negative[of "1/k" z] by auto qed lemma arg_mult_eq: assumes "z * z1 \ 0" and "z * z2 \ 0" assumes "arg (z * z1) = arg (z * z2)" shows "arg z1 = arg z2" -proof- - from assms have "\arg z + arg z1\ = \arg z + arg z2\" - by (simp add: arg_mult) - then obtain x::int where *: "arg z1 - arg z2 = 2 * x * pi" - using canon_ang_eqE[of "arg z + arg z1" "arg z + arg z2"] - by auto - moreover - have "arg z1 - arg z2 < 2*pi" "arg z1 - arg z2 > -2*pi" - using arg_bounded[of z1] arg_bounded[of z2] - by auto - ultimately - have "-1 < x" "x < 1" - using divide_strict_right_mono[of "-pi" "pi * x" pi] - by auto - hence "x = 0" - by auto - thus ?thesis - using * - by simp -qed + by (metis (no_types, lifting) arg_cis assms canon_ang_arg cis_arg mult_eq_0_iff nonzero_mult_div_cancel_left sgn_divide) text \Argument of conjugate\ lemma arg_cnj_pi: assumes "arg z = pi" shows "arg (cnj z) = pi" -proof- - have "cos (arg (cnj z)) = cos (arg z)" - using rcis_cmod_arg[of z, symmetric] Re_rcis[of "cmod z" "arg z"] - using rcis_cmod_arg[of "cnj z", symmetric] Re_rcis[of "cmod (cnj z)" "arg (cnj z)"] - by auto - hence "arg (cnj z) = arg z \ arg(cnj z) = -arg z" - using arg_bounded[of z] arg_bounded[of "cnj z"] - by (metis arccos_cos arccos_cos2 less_eq_real_def linorder_le_cases minus_minus) - thus ?thesis - using assms - using arg_bounded[of "cnj z"] - by auto -qed + using arg_pi_iff assms by auto lemma arg_cnj_not_pi: assumes "arg z \ pi" shows "arg (cnj z) = -arg z" proof(cases "arg z = 0") case True thus ?thesis using eq_cnj_iff_real[of z] is_real_arg1[of z] by force next case False - have "cos (arg (cnj z)) = cos (arg z)" - using rcis_cmod_arg[of z] Re_rcis[of "cmod z" "arg z"] - using rcis_cmod_arg[of "cnj z"] Re_rcis[of "cmod (cnj z)" "arg (cnj z)"] - by auto - hence "arg (cnj z) = arg z \ arg(cnj z) = -arg z" + have "arg (cnj z) = arg z \ arg(cnj z) = -arg z" using arg_bounded[of z] arg_bounded[of "cnj z"] - by (metis arccos_cos arccos_cos2 less_eq_real_def linorder_le_cases minus_minus) + by (smt (verit, best) arccos_cos arccos_cos2 cnj.sel(1) complex_cnj_zero_iff complex_mod_cnj cos_arg) moreover - have "sin (arg (cnj z)) = -sin (arg z)" - using rcis_cmod_arg[of z] Im_rcis[of "cmod z" "arg z"] - using rcis_cmod_arg[of "cnj z"] Im_rcis[of "cmod (cnj z)" "arg (cnj z)"] - using calculation eq_cnj_iff_real is_real_arg2 - by force - hence "arg (cnj z) \ arg z" + have "arg (cnj z) \ arg z" using sin_0_iff_canon[of "arg (cnj z)"] arg_bounded False assms - by auto + by (metis complex_mod_cnj eq_cnj_iff_real is_real_arg2 rcis_cmod_arg) ultimately show ?thesis by auto qed text \Argument of reciprocal\ lemma arg_inv_not_pi: assumes "z \ 0" and "arg z \ pi" shows "arg (1 / z) = - arg z" proof- have "1/z = cnj z / cor ((cmod z)\<^sup>2 )" using \z \ 0\ complex_mult_cnj_cmod[of z] by (auto simp add:field_simps) thus ?thesis using arg_div_real_positive[of "(cmod z)\<^sup>2" "cnj z"] \z \ 0\ using arg_cnj_not_pi[of z] \arg z \ pi\ by auto qed lemma arg_inv_pi: assumes "z \ 0" and "arg z = pi" shows "arg (1 / z) = pi" proof- have "1/z = cnj z / cor ((cmod z)\<^sup>2 )" using \z \ 0\ complex_mult_cnj_cmod[of z] by (auto simp add:field_simps) thus ?thesis using arg_div_real_positive[of "(cmod z)\<^sup>2" "cnj z"] \z \ 0\ using arg_cnj_pi[of z] \arg z = pi\ by auto qed lemma arg_inv_2kpi: assumes "z \ 0" shows "\ k::int. arg (1 / z) = - arg z + 2*k*pi" using arg_inv_pi[OF assms] using arg_inv_not_pi[OF assms] by (cases "arg z = pi") (rule_tac x="1" in exI, simp, rule_tac x="0" in exI, simp) lemma arg_inv: assumes "z \ 0" shows "arg (1 / z) = \- arg z\" -proof- - obtain k::int where "arg(1 / z) = - arg z + 2*k*pi" - using arg_inv_2kpi[of z] - using assms - by auto - hence "\arg(1 / z)\ = \- arg z\" - using canon_ang_eq - by(simp add:field_simps) - thus ?thesis - using canon_ang_arg[of "1 / z"] - by auto -qed + by (metis arg_inv_not_pi arg_inv_pi assms canon_ang_arg canon_ang_uminus_pi) text \Argument of quotient\ lemma arg_div_2kpi: assumes "z1 \ 0" and "z2 \ 0" shows "\ k::int. arg (z1 / z2) = arg z1 - arg z2 + 2*k*pi" proof- obtain x1 where "arg (z1 * (1 / z2)) = arg z1 + arg (1 / z2) + 2 * real_of_int x1 * pi" using assms arg_mult_2kpi[of z1 "1/z2"] by auto moreover obtain x2 where "arg (1 / z2) = - arg z2 + 2 * real_of_int x2 * pi" using assms arg_inv_2kpi[of z2] by auto ultimately show ?thesis by (rule_tac x="x1 + x2" in exI, simp add: field_simps) qed lemma arg_div: assumes "z1 \ 0" and "z2 \ 0" shows "arg(z1 / z2) = \arg z1 - arg z2\" proof- obtain k::int where "arg(z1 / z2) = arg z1 - arg z2 + 2*k*pi" using arg_div_2kpi[of z1 z2] using assms by auto hence "canon_ang(arg(z1 / z2)) = canon_ang(arg z1 - arg z2)" using canon_ang_eq by(simp add:field_simps) thus ?thesis using canon_ang_arg[of "z1/z2"] by auto qed text \Argument of opposite\ lemma arg_uminus: assumes "z \ 0" shows "arg (-z) = \arg z + pi\" using assms using arg_mult[of "-1" z] using arg_complex_of_real_negative[of "-1"] by (auto simp add: field_simps) lemma arg_uminus_opposite_sign: assumes "z \ 0" shows "arg z > 0 \ \ arg (-z) > 0" proof (cases "arg z = 0") case True thus ?thesis using assms by (simp add: arg_uminus) next case False show ?thesis proof (cases "arg z > 0") case True thus ?thesis using assms using arg_bounded[of z] using canon_ang_plus_pi1[of "arg z"] by (simp add: arg_uminus) next case False thus ?thesis using \arg z \ 0\ using assms using arg_bounded[of z] using canon_ang_plus_pi2[of "arg z"] by (simp add: arg_uminus) qed qed text \Sign of argument is the same as the sign of the Imaginary part\ lemma arg_Im_sgn: assumes "\ is_real z" shows "sgn (arg z) = sgn (Im z)" proof- have "z \ 0" using assms by auto then obtain r \ where polar: "z = cor r * cis \" "\ = arg z" "r > 0" by (smt cmod_cis mult_eq_0_iff norm_ge_zero of_real_0) hence "Im z = r * sin \" by (metis Im_mult_real Re_complex_of_real cis.simps(2) Im_complex_of_real) hence "Im z > 0 \ sin \ > 0" "Im z < 0 \ sin \ < 0" using \r > 0\ using mult_pos_pos mult_nonneg_nonneg zero_less_mult_pos mult_less_cancel_left by smt+ moreover have "\ \ pi" "\ \ 0" using \\ is_real z\ polar cis_pi by force+ hence "sin \ > 0 \ \ > 0" "\ < 0 \ sin \ < 0" using \\ = arg z\ \\ \ 0\ \\ \ pi\ using arg_bounded[of z] by (smt sin_gt_zero sin_le_zero sin_pi_minus sin_0_iff_canon sin_ge_zero)+ ultimately show ?thesis using \\ = arg z\ by auto qed subsubsection \Complex square root\ definition "ccsqrt z = rcis (sqrt (cmod z)) (arg z / 2)" lemma square_ccsqrt [simp]: shows "(ccsqrt x)\<^sup>2 = x" unfolding ccsqrt_def by (subst DeMoivre2) (simp add: rcis_cmod_arg) lemma ex_complex_sqrt: shows "\ s::complex. s*s = z" unfolding power2_eq_square[symmetric] by (rule_tac x="csqrt z" in exI) simp lemma ccsqrt: assumes "s * s = z" shows "s = ccsqrt z \ s = -ccsqrt z" proof (cases "s = 0") case True thus ?thesis using assms unfolding ccsqrt_def by simp next case False then obtain k::int where "cmod s * cmod s = cmod z" "2 * arg s - arg z = 2*k*pi" using assms using rcis_cmod_arg[of z] rcis_cmod_arg[of s] using arg_mult[of s s] using canon_ang(3)[of "2*arg s"] by (auto simp add: norm_mult arg_mult) have *: "sqrt (cmod z) = cmod s" using \cmod s * cmod s = cmod z\ by (smt norm_not_less_zero real_sqrt_abs2) have **: "arg z / 2 = arg s - k*pi" using \2 * arg s - arg z = 2*k*pi\ by simp have "cis (arg s - k*pi) = cis (arg s) \ cis (arg s - k*pi) = -cis (arg s)" proof (cases "even k") case True hence "cis (arg s - k*pi) = cis (arg s)" by (simp add: cis_def complex.corec cos_diff sin_diff) thus ?thesis by simp next case False hence "cis (arg s - k*pi) = -cis (arg s)" by (simp add: cis_def complex.corec Complex_eq cos_diff sin_diff) thus ?thesis by simp qed thus ?thesis proof assume ***: "cis (arg s - k * pi) = cis (arg s)" hence "s = ccsqrt z" using rcis_cmod_arg[of s] unfolding ccsqrt_def rcis_def by (subst *, subst **, subst ***, simp) thus ?thesis by simp next assume ***: "cis (arg s - k * pi) = -cis (arg s)" hence "s = - ccsqrt z" using rcis_cmod_arg[of s] unfolding ccsqrt_def rcis_def by (subst *, subst **, subst ***, simp) thus ?thesis by simp qed qed lemma null_ccsqrt [simp]: shows "ccsqrt x = 0 \ x = 0" unfolding ccsqrt_def by auto lemma ccsqrt_mult: shows "ccsqrt (a * b) = ccsqrt a * ccsqrt b \ ccsqrt (a * b) = - ccsqrt a * ccsqrt b" proof (cases "a = 0 \ b = 0") case True thus ?thesis by auto next case False obtain k::int where "arg a + arg b - \arg a + arg b\ = 2 * real_of_int k * pi" using canon_ang(3)[of "arg a + arg b"] by auto hence *: "\arg a + arg b\ = arg a + arg b - 2 * (real_of_int k) * pi" by (auto simp add: field_simps) have "cis (\arg a + arg b\ / 2) = cis (arg a / 2 + arg b / 2) \ cis (\arg a + arg b\ / 2) = - cis (arg a / 2 + arg b / 2)" using cos_even_kpi[of k] cos_odd_kpi[of k] by ((subst *)+, (subst diff_divide_distrib)+, (subst add_divide_distrib)+) (cases "even k", auto simp add: cis_def complex.corec Complex_eq cos_diff sin_diff) thus ?thesis using False unfolding ccsqrt_def - by (simp add: rcis_mult real_sqrt_mult arg_mult) - (auto simp add: rcis_def) + by (smt (verit, best) arg_mult mult_minus_left mult_minus_right no_zero_divisors norm_mult rcis_def rcis_mult real_sqrt_mult) qed lemma csqrt_real: assumes "is_real x" shows "(Re x \ 0 \ ccsqrt x = cor (sqrt (Re x))) \ (Re x < 0 \ ccsqrt x = \ * cor (sqrt (- (Re x))))" proof (cases "x = 0") case True thus ?thesis by auto next case False show ?thesis proof (cases "Re x > 0") case True hence "arg x = 0" using \is_real x\ by (metis arg_complex_of_real_positive complex_of_real_Re) thus ?thesis using \Re x > 0\ \is_real x\ unfolding ccsqrt_def by (simp add: cmod_eq_Re) next case False hence "Re x < 0" using \x \ 0\ \is_real x\ using complex_eq_if_Re_eq by auto hence "arg x = pi" using \is_real x\ by (metis arg_complex_of_real_negative complex_of_real_Re) thus ?thesis using \Re x < 0\ \is_real x\ unfolding ccsqrt_def rcis_def by (simp add: cis_def complex.corec Complex_eq cmod_eq_Re) qed qed text \Rotation of complex vector to x-axis.\ lemma is_real_rot_to_x_axis: assumes "z \ 0" shows "is_real (cis (-arg z) * z)" proof (cases "arg z = pi") case True thus ?thesis using is_real_arg1[of z] by auto next case False hence "\- arg z\ = - arg z" using canon_ang_eqI[of "- arg z" "-arg z"] using arg_bounded[of z] by (auto simp add: field_simps) hence "arg (cis (- (arg z)) * z) = 0" using arg_mult[of "cis (- (arg z))" z] \z \ 0\ using arg_cis[of "- arg z"] by simp thus ?thesis using is_real_arg1[of "cis (- arg z) * z"] by auto qed lemma positive_rot_to_x_axis: assumes "z \ 0" shows "Re (cis (-arg z) * z) > 0" using assms by (smt Re_complex_of_real cis_rcis_eq mult_cancel_right1 rcis_cmod_arg rcis_mult rcis_zero_arg zero_less_norm_iff) text \Inequalities involving @{term cmod}.\ lemma cmod_1_plus_mult_le: shows "cmod (1 + z*w) \ sqrt((1 + (cmod z)\<^sup>2) * (1 + (cmod w)\<^sup>2))" proof- have "Re ((1+z*w)*(1+cnj z*cnj w)) \ Re (1+z*cnj z)* Re (1+w*cnj w)" proof- have "Re ((w - cnj z)*cnj(w - cnj z)) \ 0" by (subst complex_mult_cnj_cmod) (simp add: power2_eq_square) hence "Re (z*w + cnj z * cnj w) \ Re (w*cnj w) + Re(z*cnj z)" by (simp add: field_simps) thus ?thesis by (simp add: field_simps) qed hence "(cmod (1 + z * w))\<^sup>2 \ (1 + (cmod z)\<^sup>2) * (1 + (cmod w)\<^sup>2)" by (subst cmod_square)+ simp thus ?thesis by (metis abs_norm_cancel real_sqrt_abs real_sqrt_le_iff) qed lemma cmod_diff_ge: shows "cmod (b - c) \ sqrt (1 + (cmod b)\<^sup>2) - sqrt (1 + (cmod c)\<^sup>2)" proof- have "(cmod (b - c))\<^sup>2 + (1/2*Im(b*cnj c - c*cnj b))\<^sup>2 \ 0" by simp hence "(cmod (b - c))\<^sup>2 \ - (1/2*Im(b*cnj c - c*cnj b))\<^sup>2" by simp hence "(cmod (b - c))\<^sup>2 \ (1/2*Re(b*cnj c + c*cnj b))\<^sup>2 - Re(b*cnj b*c*cnj c) " by (auto simp add: power2_eq_square field_simps) hence "Re ((b - c)*(cnj b - cnj c)) \ (1/2*Re(b*cnj c + c*cnj b))\<^sup>2 - Re(b*cnj b*c*cnj c)" by (subst (asm) cmod_square) simp moreover have "(1 + (cmod b)\<^sup>2) * (1 + (cmod c)\<^sup>2) = 1 + Re(b*cnj b) + Re(c*cnj c) + Re(b*cnj b*c*cnj c)" by (subst cmod_square)+ (simp add: field_simps power2_eq_square) moreover have "(1 + Re (scalprod b c))\<^sup>2 = 1 + 2*Re(scalprod b c) + ((Re (scalprod b c))\<^sup>2)" by (subst power2_sum) simp hence "(1 + Re (scalprod b c))\<^sup>2 = 1 + Re(b*cnj c + c*cnj b) + (1/2 * Re (b*cnj c + c*cnj b))\<^sup>2" by simp ultimately have "(1 + (cmod b)\<^sup>2) * (1 + (cmod c)\<^sup>2) \ (1 + Re (scalprod b c))\<^sup>2" by (simp add: field_simps) moreover have "sqrt((1 + (cmod b)\<^sup>2) * (1 + (cmod c)\<^sup>2)) \ 0" by (metis one_power2 real_sqrt_sum_squares_mult_ge_zero) ultimately have "sqrt((1 + (cmod b)\<^sup>2) * (1 + (cmod c)\<^sup>2)) \ 1 + Re (scalprod b c)" by (metis power2_le_imp_le real_sqrt_ge_0_iff real_sqrt_pow2_iff) hence "Re ((b - c) * (cnj b - cnj c)) \ 1 + Re (c*cnj c) + 1 + Re (b*cnj b) - 2*sqrt((1 + (cmod b)\<^sup>2) * (1 + (cmod c)\<^sup>2))" by (simp add: field_simps) hence *: "(cmod (b - c))\<^sup>2 \ (sqrt (1 + (cmod b)\<^sup>2) - sqrt (1 + (cmod c)\<^sup>2))\<^sup>2" apply (subst cmod_square)+ apply (subst (asm) cmod_square)+ apply (subst power2_diff) apply (subst real_sqrt_pow2, simp) apply (subst real_sqrt_pow2, simp) apply (simp add: real_sqrt_mult) done thus ?thesis proof (cases "sqrt (1 + (cmod b)\<^sup>2) - sqrt (1 + (cmod c)\<^sup>2) > 0") case True thus ?thesis using power2_le_imp_le[OF *] by simp next case False hence "0 \ sqrt (1 + (cmod b)\<^sup>2) - sqrt (1 + (cmod c)\<^sup>2)" by (metis less_eq_real_def linorder_neqE_linordered_idom) moreover have "cmod (b - c) \ 0" by simp ultimately show ?thesis by (metis add_increasing monoid_add_class.add.right_neutral) qed qed lemma cmod_diff_le: shows "cmod (b - c) \ sqrt (1 + (cmod b)\<^sup>2) + sqrt (1 + (cmod c)\<^sup>2)" proof- have "(cmod (b + c))\<^sup>2 + (1/2*Im(b*cnj c - c*cnj b))\<^sup>2 \ 0" by simp hence "(cmod (b + c))\<^sup>2 \ - (1/2*Im(b*cnj c - c*cnj b))\<^sup>2" by simp hence "(cmod (b + c))\<^sup>2 \ (1/2*Re(b*cnj c + c*cnj b))\<^sup>2 - Re(b*cnj b*c*cnj c) " by (auto simp add: power2_eq_square field_simps) hence "Re ((b + c)*(cnj b + cnj c)) \ (1/2*Re(b*cnj c + c*cnj b))\<^sup>2 - Re(b*cnj b*c*cnj c)" by (subst (asm) cmod_square) simp moreover have "(1 + (cmod b)\<^sup>2) * (1 + (cmod c)\<^sup>2) = 1 + Re(b*cnj b) + Re(c*cnj c) + Re(b*cnj b*c*cnj c)" by (subst cmod_square)+ (simp add: field_simps power2_eq_square) moreover have ++: "2*Re(scalprod b c) = Re(b*cnj c + c*cnj b)" by simp have "(1 - Re (scalprod b c))\<^sup>2 = 1 - 2*Re(scalprod b c) + ((Re (scalprod b c))\<^sup>2)" by (subst power2_diff) simp hence "(1 - Re (scalprod b c))\<^sup>2 = 1 - Re(b*cnj c + c*cnj b) + (1/2 * Re (b*cnj c + c*cnj b))\<^sup>2" by (subst ++[symmetric]) simp ultimately have "(1 + (cmod b)\<^sup>2) * (1 + (cmod c)\<^sup>2) \ (1 - Re (scalprod b c))\<^sup>2" by (simp add: field_simps) moreover have "sqrt((1 + (cmod b)\<^sup>2) * (1 + (cmod c)\<^sup>2)) \ 0" by (metis one_power2 real_sqrt_sum_squares_mult_ge_zero) ultimately have "sqrt((1 + (cmod b)\<^sup>2) * (1 + (cmod c)\<^sup>2)) \ 1 - Re (scalprod b c)" by (metis power2_le_imp_le real_sqrt_ge_0_iff real_sqrt_pow2_iff) hence "Re ((b - c) * (cnj b - cnj c)) \ 1 + Re (c*cnj c) + 1 + Re (b*cnj b) + 2*sqrt((1 + (cmod b)\<^sup>2) * (1 + (cmod c)\<^sup>2))" by (simp add: field_simps) hence *: "(cmod (b - c))\<^sup>2 \ (sqrt (1 + (cmod b)\<^sup>2) + sqrt (1 + (cmod c)\<^sup>2))\<^sup>2" apply (subst cmod_square)+ apply (subst (asm) cmod_square)+ apply (subst power2_sum) apply (subst real_sqrt_pow2, simp) apply (subst real_sqrt_pow2, simp) apply (simp add: real_sqrt_mult) done thus ?thesis using power2_le_imp_le[OF *] by simp qed text \Definition of Euclidean distance between two complex numbers.\ definition cdist where [simp]: "cdist z1 z2 \ cmod (z2 - z1)" text \Misc. properties of complex numbers.\ lemma ex_complex_to_complex [simp]: fixes z1 z2 :: complex assumes "z1 \ 0" and "z2 \ 0" shows "\k. k \ 0 \ z2 = k * z1" using assms by (rule_tac x="z2/z1" in exI) simp lemma ex_complex_to_one [simp]: fixes z::complex assumes "z \ 0" shows "\k. k \ 0 \ k * z = 1" using assms by (rule_tac x="1/z" in exI) simp lemma ex_complex_to_complex2 [simp]: fixes z::complex shows "\k. k \ 0 \ k * z = z" by (rule_tac x="1" in exI) simp lemma complex_sqrt_1: fixes z::complex assumes "z \ 0" shows "z = 1 / z \ z = 1 \ z = -1" using assms using nonzero_eq_divide_eq square_eq_iff by fastforce end diff --git a/thys/Complex_Geometry/More_Transcendental.thy b/thys/Complex_Geometry/More_Transcendental.thy --- a/thys/Complex_Geometry/More_Transcendental.thy +++ b/thys/Complex_Geometry/More_Transcendental.thy @@ -1,409 +1,224 @@ (* ---------------------------------------------------------------------------- *) section \Introduction\ (* ---------------------------------------------------------------------------- *) text \The complex plane or some of its parts (e.g., the unit disc or the upper half plane) are often taken as the domain in which models of various geometries (both Euclidean and non-Euclidean ones) are formalized. The complex plane gives simpler and more compact formulas than the Cartesian plane. Within complex plane is easier to describe geometric objects and perform the calculations (usually shedding some new light on the subject). We give a formalization of the extended complex plane (given both as a complex projective space and as the Riemann sphere), its objects (points, circles and lines), and its transformations (Möbius transformations).\ (* ---------------------------------------------------------------------------- *) section \Related work\ (* ---------------------------------------------------------------------------- *) text\During the last decade, there have been many results in formalizing geometry in proof-assistants. Parts of Hilbert’s seminal book ,,Foundations of Geometry'' \cite{hilbert} have been formalized both in Coq and Isabelle/Isar. Formalization of first two groups of axioms in Coq, in an intuitionistic setting was done by Dehlinger et al. \cite{hilbert-coq}. First formalization in Isabelle/HOL was done by Fleuriot and Meikele \cite{hilbert-isabelle}, and some further developments were made in master thesis of Scott \cite{hilbert-scott}. Large fragments of Tarski's geometry \cite{tarski} have been formalized in Coq by Narboux et al. \cite{narboux-tarski}. Within Coq, there are also formalizations of von Plato’s constructive geometry by Kahn \cite{vonPlato,von-plato-formalization}, French high school geometry by Guilhot \cite{guilhot} and ruler and compass geometry by Duprat \cite{duprat2008}, etc. In our previous work \cite{petrovic2012formalizing}, we have already formally investigated a Cartesian model of Euclidean geometry. \ (* ---------------------------------------------------------------------------- *) section \Background theories\ (* ---------------------------------------------------------------------------- *) text \In this section we introduce some basic mathematical notions and prove some lemmas needed in the rest of our formalization. We describe: \<^item> trigonometric functions, \<^item> complex numbers, \<^item> systems of two and three linear equations with two unknowns (over arbitrary fields), \<^item> quadratic equations (over real and complex numbers), systems of quadratic and real equations, and systems of two quadratic equations, \<^item> two-dimensional vectors and matrices over complex numbers. \ (* -------------------------------------------------------------------------- *) subsection \Library Additions for Trigonometric Functions\ (* -------------------------------------------------------------------------- *) theory More_Transcendental - imports Complex_Main + imports Complex_Main "HOL-Library.Periodic_Fun" begin text \Additional properties of @{term sin} and @{term cos} functions that are later used in proving conjectures for argument of complex number.\ text \Sign of trigonometric functions on some characteristic intervals.\ lemma cos_lt_zero_on_pi2_pi [simp]: assumes "x > pi/2" and "x \ pi" shows "cos x < 0" using cos_gt_zero_pi[of "pi - x"] assms by simp text \Value of trigonometric functions in points $k\pi$ and $\frac{\pi}{2} + k\pi$.\ lemma sin_kpi [simp]: fixes k::int shows "sin (k * pi) = 0" by (simp add: sin_zero_iff_int2) lemma cos_odd_kpi [simp]: fixes k::int assumes "odd k" shows "cos (k * pi) = -1" -proof (cases "k \ 0") - case True - hence "odd (nat k)" - using \odd k\ - by (auto simp add: even_nat_iff) - thus ?thesis - using \k \ 0\ cos_npi[of "nat k"] - by auto -next - case False - hence "-k \ 0" "odd (nat (-k))" - using \odd k\ - by (auto simp add: even_nat_iff) - thus ?thesis - using cos_npi[of "nat (-k)"] - by auto -qed + by (simp add: assms mult.commute) lemma cos_even_kpi [simp]: fixes k::int assumes "even k" shows "cos (k * pi) = 1" -proof (cases "k \ 0") - case True - hence "even (nat k)" - using \even k\ - by (simp add: even_nat_iff) - thus ?thesis - using \k \ 0\ cos_npi[of "nat k"] - by auto -next - case False - hence "-k \ 0" "even (nat (-k))" - using \even k\ - by (auto simp add: even_nat_iff) - thus ?thesis - using cos_npi[of "nat (-k)"] - by auto -qed + by (simp add: assms mult.commute) lemma sin_pi2_plus_odd_kpi [simp]: fixes k::int assumes "odd k" shows "sin (pi / 2 + k * pi) = -1" using assms by (simp add: sin_add) lemma sin_pi2_plus_even_kpi [simp]: fixes k::int assumes "even k" shows "sin (pi / 2 + k * pi) = 1" using assms by (simp add: sin_add) text \Solving trigonometric equations and systems with special values (0, 1, or -1) of sine and cosine functions\ lemma cos_0_iff_canon: assumes "cos \ = 0" and "-pi < \" and "\ \ pi" shows "\ = pi/2 \ \ = -pi/2" -proof- - obtain k::int where "odd k" "\ = k * pi/2" - using cos_zero_iff_int[of \] assms(1) - by auto - thus ?thesis - proof (cases "k > 1 \ k < -1") - case True - hence "k \ 3 \ k \ -3" - using \odd k\ - by (smt dvd_refl even_minus) - hence "\ \ 3*pi/2 \ \ \ -3*pi/2" - using mult_right_mono[of k "-3" "pi / 2"] - using \\ = k * pi/2\ - by auto - thus ?thesis - using \- pi < \\ \\ \ pi\ - by auto - next - case False - hence "k = -1 \ k = 0 \ k = 1" - by auto - hence "k = -1 \ k = 1" - using \odd k\ - by auto - thus ?thesis - using \\ = k * pi/2\ - by auto - qed -qed + by (smt (verit, best) arccos_0 arccos_cos assms cos_minus divide_minus_left) lemma sin_0_iff_canon: assumes "sin \ = 0" and "-pi < \" and "\ \ pi" shows "\ = 0 \ \ = pi" -proof- - obtain k::int where "even k" "\ = k * pi/2" - using sin_zero_iff_int[of \] assms(1) - by auto - thus ?thesis - proof (cases "k > 2 \ k < 0") - case True - hence "k \ 4 \ k \ -2" - using \even k\ - by (smt evenE) - hence "\ \ 2*pi \ \ \ -pi" - proof - assume "4 \ k" - hence "4 * pi/2 \ \" - using mult_right_mono[of "4" "k" "pi/2"] - by (subst \\ = k * pi/2\) auto - thus ?thesis - by simp - next - assume "k \ -2" - hence "-2*pi/2 \ \" - using mult_right_mono[of "k" "-2" "pi/2"] - by (subst \\ = k * pi/2\, auto) - thus ?thesis - by simp - qed - thus ?thesis - using \- pi < \\ \\ \ pi\ - by auto - next - case False - hence "k = 0 \ k = 1 \ k = 2" - by auto - hence "k = 0 \ k = 2" - using \even k\ - by auto - thus ?thesis - using \\ = k * pi/2\ - by auto - qed -qed + using assms sin_eq_0_pi by force lemma cos0_sin1: - assumes "cos \ = 0" and "sin \ = 1" + assumes "sin \ = 1" shows "\ k::int. \ = pi/2 + 2*k*pi" -proof- - from \cos \ = 0\ - obtain k::int where "odd k" "\ = k * (pi / 2)" - using cos_zero_iff_int[of "\"] - by auto - then obtain k'::int where "k = 2*k' + 1" - using oddE by blast - hence "\ = pi/2 + (k' * pi)" - using \\ = k * (pi / 2)\ - by (auto simp add: field_simps) - hence "even k'" - using \sin \ = 1\ sin_pi2_plus_odd_kpi[of k'] - by auto - thus ?thesis - using \\ = pi /2 + (k' * pi)\ - unfolding even_iff_mod_2_eq_zero - by auto -qed - -lemma cos1_sin0: - assumes "cos \ = 1" and "sin \ = 0" - shows "\ k::int. \ = 2*k*pi" -proof- - from \sin \ = 0\ - obtain k::int where "even k" "\ = k * (pi / 2)" - using sin_zero_iff_int[of "\"] - by auto - then obtain k'::int where "k = 2*k'" - using evenE by blast - hence "\ = k' * pi" - using \\ = k * (pi / 2)\ - by (auto simp add: field_simps) - hence "even k'" - using \cos \ = 1\ cos_odd_kpi[of k'] - by auto - thus ?thesis - using \\ = k' * pi\ - using assms(1) cos_one_2pi_int by auto -qed + by (smt (verit, ccfv_threshold) assms cos_diff cos_one_2pi_int cos_pi_half mult_cancel_right1 sin_pi_half sin_plus_pi) (* TODO: add lemmas for cos = -1, sin = 0 and cos = 0, sin = -1 *) text \Sine is injective on $[-\frac{\pi}{2}, \frac{\pi}{2}]$\ lemma sin_inj: assumes "-pi/2 \ \ \ \ \ pi/2" and "-pi/2 \ \' \ \' \ pi/2" assumes "\ \ \'" shows "sin \ \ sin \'" - using assms - using sin_monotone_2pi[of \ \'] sin_monotone_2pi[of \' \] - by (cases "\ < \'") auto + by (metis assms divide_minus_left sin_inj_pi) text \Periodicity of trigonometric functions\ text \The following are available in HOL-Decision\_Procs.Approximation\_Bounds, but we want to avoid that dependency\ lemma sin_periodic_nat [simp]: fixes n :: nat shows "sin (x + n * (2 * pi)) = sin x" -proof (induct n arbitrary: x) - case (Suc n) - have split_pi_off: "x + (Suc n) * (2 * pi) = (x + n * (2 * pi)) + 2 * pi" - unfolding Suc_eq_plus1 distrib_right - by (auto simp add: field_simps) - show ?case unfolding split_pi_off using Suc by auto -qed auto + by (metis (no_types, hide_lams) add.commute add.left_neutral cos_2npi cos_one_2pi_int mult.assoc mult.commute mult.left_neutral mult_zero_left sin_add sin_int_2pin) lemma sin_periodic_int [simp]: fixes i :: int shows "sin (x + i * (2 * pi)) = sin x" -proof(cases "0 \ i") - case True - thus ?thesis - using sin_periodic_nat[of x "nat i"] - by auto -next - case False hence i_nat: "i = - real (nat (-i))" by auto - have "sin x = sin (x + i * (2 * pi) - i * (2 * pi))" by auto - also have "\ = sin (x + i * (2 * pi))" - unfolding i_nat mult_minus_left diff_minus_eq_add by (rule sin_periodic_nat) - finally show ?thesis by auto -qed + by (metis add.right_neutral cos_int_2pin mult.commute mult.right_neutral mult_zero_right sin_add sin_int_2pin) lemma cos_periodic_nat [simp]: fixes n :: nat shows "cos (x + n * (2 * pi)) = cos x" -proof (induct n arbitrary: x) - case (Suc n) - have split_pi_off: "x + (Suc n) * (2 * pi) = (x + n * (2 * pi)) + 2 * pi" - unfolding Suc_eq_plus1 distrib_right - by (auto simp add: field_simps) - show ?case unfolding split_pi_off using Suc by auto -qed auto + by (metis add.left_neutral cos_2npi cos_add cos_periodic mult.assoc mult_2 mult_2_right of_nat_numeral sin_periodic sin_periodic_nat) lemma cos_periodic_int [simp]: fixes i :: int shows "cos (x + i * (2 * pi)) = cos x" -proof(cases "0 \ i") - case True - thus ?thesis - using cos_periodic_nat[of x "nat i"] - by auto -next - case False hence i_nat: "i = - real (nat (-i))" by auto - have "cos x = cos (x + i * (2 * pi) - i * (2 * pi))" by auto - also have "\ = cos (x + i * (2 * pi))" - unfolding i_nat mult_minus_left diff_minus_eq_add by (rule cos_periodic_nat) - finally show ?thesis by auto -qed + by (metis cos_add cos_int_2pin diff_zero mult.commute mult.right_neutral mult_zero_right sin_int_2pin) text \Values of both sine and cosine are repeated only after multiples of $2\cdot \pi$\ lemma sin_cos_eq: fixes a b :: real assumes "cos a = cos b" and "sin a = sin b" shows "\ k::int. a - b = 2*k*pi" -proof- - from assms have "sin (a - b) = 0" "cos (a - b) = 1" - using sin_diff[of a b] cos_diff[of a b] - by auto - thus ?thesis - using cos1_sin0 - by auto -qed + by (metis assms cos_diff cos_one_2pi_int mult.commute sin_cos_squared_add3) text \The following two lemmas are consequences of surjectivity of cosine for the range $[-1, 1]$.\ lemma ex_cos_eq: assumes "-pi/2 \ \ \ \ \ pi/2" assumes "a \ 0" and "a < 1" shows "\ \'. -pi/2 \ \' \ \' \ pi/2 \ \' \ \ \ cos (\ - \') = a" proof- have "arccos a > 0" "arccos a \ pi/2" using \a \ 0\ \a < 1\ using arccos_lt_bounded arccos_le_pi2 by auto show ?thesis proof (cases "\ - arccos a \ - pi/2") case True thus ?thesis using assms \arccos a > 0\ \arccos a \ pi/2\ by (rule_tac x = "\ - arccos a" in exI) auto next case False thus ?thesis using assms \arccos a > 0\ \arccos a \ pi/2\ by (rule_tac x = "\ + arccos a" in exI) auto qed qed lemma ex_cos_gt: assumes "-pi/2 \ \ \ \ \ pi/2" assumes "a < 1" shows "\ \'. -pi/2 \ \' \ \' \ pi/2 \ \' \ \ \ cos (\ - \') > a" proof- - have "\ a'. a' \ 0 \ a' > a \ a' < 1" - using \a < 1\ - using divide_strict_right_mono[of "2*a + (1 - a)" 2 2] - by (rule_tac x="if a < 0 then 0 else a + (1-a)/2" in exI) (auto simp add: field_simps) - then obtain a' where "a' \ 0" "a' > a" "a' < 1" - by auto + obtain a' where "a' \ 0" "a' > a" "a' < 1" + by (metis assms(2) dense_le_bounded linear not_one_le_zero) thus ?thesis using ex_cos_eq[of \ a'] assms by auto qed text \The function @{term atan2} is a generalization of @{term arctan} that takes a pair of coordinates of non-zero points returns its angle in the range $[-\pi, \pi)$.\ definition atan2 where "atan2 y x = (if x > 0 then arctan (y/x) else if x < 0 then if y > 0 then arctan (y/x) + pi else arctan (y/x) - pi else if y > 0 then pi/2 else if y < 0 then -pi/2 else 0)" lemma atan2_bounded: shows "-pi \ atan2 y x \ atan2 y x < pi" using arctan_bounded[of "y/x"] zero_le_arctan_iff[of "y/x"] arctan_le_zero_iff[of "y/x"] zero_less_arctan_iff[of "y/x"] arctan_less_zero_iff[of "y/x"] using divide_neg_neg[of y x] divide_neg_pos[of y x] divide_pos_pos[of y x] divide_pos_neg[of y x] unfolding atan2_def by (simp (no_asm_simp)) auto end diff --git a/thys/Complex_Geometry/Riemann_Sphere.thy b/thys/Complex_Geometry/Riemann_Sphere.thy --- a/thys/Complex_Geometry/Riemann_Sphere.thy +++ b/thys/Complex_Geometry/Riemann_Sphere.thy @@ -1,641 +1,641 @@ (* ---------------------------------------------------------------------------- *) section \Riemann sphere\ (* ---------------------------------------------------------------------------- *) text \The extended complex plane $\mathbb{C}P^1$ can be identified with a Riemann (unit) sphere $\Sigma$ by means of stereographic projection. The sphere is projected from its north pole $N$ to the $xOy$ plane (identified with $\mathbb{C}$). This projection establishes a bijective map $sp$ between $\Sigma \setminus \{N\}$ and the finite complex plane $\mathbb{C}$. The infinite point is defined as the image of $N$.\ theory Riemann_Sphere imports Homogeneous_Coordinates Circlines "HOL-Analysis.Product_Vector" begin text \Coordinates in $\mathbb{R}^3$\ type_synonym R3 = "real \ real \ real" text \Type of points of $\Sigma$\ abbreviation unit_sphere where "unit_sphere \ {(x::real, y::real, z::real). x*x + y*y + z*z = 1}" typedef riemann_sphere = "unit_sphere" by (rule_tac x="(1, 0, 0)" in exI) simp setup_lifting type_definition_riemann_sphere lemma sphere_bounds': assumes "x*x + y*y + z*z = (1::real)" shows "-1 \ x \ x \ 1" proof- from assms have "x*x \ 1" by (smt real_minus_mult_self_le) hence "x\<^sup>2 \ 1\<^sup>2" "(- x)\<^sup>2 \ 1\<^sup>2" by (auto simp add: power2_eq_square) show "-1 \ x \ x \ 1" proof (cases "x \ 0") case True thus ?thesis using \x\<^sup>2 \ 1\<^sup>2\ by (smt power2_le_imp_le) next case False thus ?thesis using \(-x)\<^sup>2 \ 1\<^sup>2\ by (smt power2_le_imp_le) qed qed lemma sphere_bounds: assumes "x*x + y*y + z*z = (1::real)" shows "-1 \ x \ x \ 1" "-1 \ y \ y \ 1" "-1 \ z \ z \ 1" using assms using sphere_bounds'[of x y z] sphere_bounds'[of y x z] sphere_bounds'[of z x y] by (auto simp add: field_simps) (* ---------------------------------------------------------------------------- *) subsection \Parametrization of the unit sphere in polar coordinates\ (* ---------------------------------------------------------------------------- *) lemma sphere_params_on_sphere: fixes \ \ :: real assumes "x = cos \ * cos \" and "y = cos \ * sin \" "z = sin \" shows "x*x + y*y + z*z = 1" proof- have "x*x + y*y = (cos \ * cos \) * (cos \ * cos \) + (cos \ * cos \) * (sin \ * sin \)" using assms by simp hence "x*x + y*y = cos \ * cos \" using sin_cos_squared_add3[of \] by (subst (asm) distrib_left[symmetric]) (simp add: field_simps) thus ?thesis using assms using sin_cos_squared_add3[of \] by simp qed lemma sphere_params: fixes x y z :: real assumes "x*x + y*y + z*z = 1" shows "x = cos (arcsin z) * cos (atan2 y x) \ y = cos (arcsin z) * sin (atan2 y x) \ z = sin (arcsin z)" proof (cases "z=1 \ z = -1") case True hence "x = 0 \ y = 0" using assms by auto thus ?thesis using \z = 1 \ z = -1\ by (auto simp add: cos_arcsin) next case False hence "x \ 0 \ y \ 0" using assms by (auto simp add: square_eq_1_iff) thus ?thesis using real_sqrt_unique[of y "1 - z*z"] using real_sqrt_unique[of "-y" "1 - z*z"] using sphere_bounds[OF assms] assms by (auto simp add: cos_arcsin cos_arctan sin_arctan power2_eq_square field_simps real_sqrt_divide atan2_def) qed lemma ex_sphere_params: assumes "x*x + y*y + z*z = 1" shows "\ \ \. x = cos \ * cos \ \ y = cos \ * sin \ \ z = sin \ \ -pi / 2 \ \ \ \ \ pi / 2 \ -pi \ \ \ \ < pi" using assms arcsin_bounded[of z] sphere_bounds[of x y z] by (rule_tac x="arcsin z" in exI, rule_tac x="atan2 y x" in exI) (simp add: sphere_params arcsin_bounded atan2_bounded) (* ----------------------------------------------------------------- *) subsection \Stereographic and inverse stereographic projection\ (* ----------------------------------------------------------------- *) text \Stereographic projection\ definition stereographic_r3_cvec :: "R3 \ complex_vec" where [simp]: "stereographic_r3_cvec M = (let (x, y, z) = M in (if (x, y, z) \ (0, 0, 1) then (x + \ * y, cor (1 - z)) else (1, 0) ))" lift_definition stereographic_r3_hcoords :: "R3 \ complex_homo_coords" is stereographic_r3_cvec by (auto split: if_split_asm simp add: cor_eq_0) lift_definition stereographic :: "riemann_sphere \ complex_homo" is stereographic_r3_hcoords done text \Inverse stereographic projection\ definition inv_stereographic_cvec_r3 :: "complex_vec \ R3" where [simp]: "inv_stereographic_cvec_r3 z = ( let (z1, z2) = z in if z2 = 0 then (0, 0, 1) else let z = z1/z2; X = Re (2*z / (1 + z*cnj z)); Y = Im (2*z / (1 + z*cnj z)); Z = ((cmod z)\<^sup>2 - 1) / (1 + (cmod z)\<^sup>2) in (X, Y, Z))" lemma Re_stereographic: shows "Re (2 * z / (1 + z * cnj z)) = 2 * Re z / (1 + (cmod z)\<^sup>2)" using one_plus_square_neq_zero by (subst complex_mult_cnj_cmod, subst Re_divide_real) (auto simp add: power2_eq_square) lemma Im_stereographic: shows "Im (2 * z / (1 + z * cnj z)) = 2 * Im z / (1 + (cmod z)\<^sup>2)" using one_plus_square_neq_zero by (subst complex_mult_cnj_cmod, subst Im_divide_real) (auto simp add: power2_eq_square) lemma inv_stereographic_on_sphere: assumes "X = Re (2*z / (1 + z*cnj z))" and "Y = Im (2*z / (1 + z*cnj z))" and "Z = ((cmod z)\<^sup>2 - 1) / (1 + (cmod z)\<^sup>2)" shows "X*X + Y*Y + Z*Z = 1" proof- have "1 + (cmod z)\<^sup>2 \ 0" by (smt power2_less_0) thus ?thesis using assms by (simp add: Re_stereographic Im_stereographic) (cases z, simp add: power2_eq_square real_sqrt_mult[symmetric] add_divide_distrib[symmetric], simp add: complex_norm power2_eq_square field_simps) qed lift_definition inv_stereographic_hcoords_r3 :: "complex_homo_coords \ R3" is inv_stereographic_cvec_r3 done lift_definition inv_stereographic :: "complex_homo \ riemann_sphere" is inv_stereographic_hcoords_r3 proof transfer fix v v' assume 1: "v \ vec_zero" "v' \ vec_zero" "v \\<^sub>v v'" obtain v1 v2 v'1 v'2 where *: "v = (v1, v2)" "v' = (v'1, v'2)" by (cases v, cases v', auto) obtain x y z where **: "inv_stereographic_cvec_r3 v = (x, y, z)" by (cases "inv_stereographic_cvec_r3 v", blast) have "inv_stereographic_cvec_r3 v \ unit_sphere" proof (cases "v2 = 0") case True thus ?thesis using * by simp next case False thus ?thesis using * ** inv_stereographic_on_sphere[of x "v1 / v2" y z] - by simp + by (simp add: norm_divide) qed moreover have "inv_stereographic_cvec_r3 v = inv_stereographic_cvec_r3 v'" using 1 * ** by (auto split: if_split if_split_asm) ultimately show "inv_stereographic_cvec_r3 v \ unit_sphere \ inv_stereographic_cvec_r3 v = inv_stereographic_cvec_r3 v'" by simp qed text \North pole\ definition North_R3 :: R3 where [simp]: "North_R3 = (0, 0, 1)" lift_definition North :: "riemann_sphere" is North_R3 by simp lemma stereographic_North: shows "stereographic x = \\<^sub>h \ x = North" by (transfer, transfer, auto split: if_split_asm) text \Stereographic and inverse stereographic projection are mutually inverse.\ lemma stereographic_inv_stereographic': assumes z: "z = z1/z2" and "z2 \ 0" and X: "X = Re (2*z / (1 + z*cnj z))" and Y: "Y = Im (2*z / (1 + z*cnj z))" and Z: "Z = ((cmod z)\<^sup>2 - 1) / (1 + (cmod z)\<^sup>2)" shows "\ k. k \ 0 \ (X + \*Y, complex_of_real (1 - Z)) = k *\<^sub>s\<^sub>v (z1, z2)" proof- have "1 + (cmod z)\<^sup>2 \ 0" by (metis one_power2 sum_power2_eq_zero_iff zero_neq_one) hence "(1 - Z) = 2 / (1 + (cmod z)\<^sup>2)" using Z by (auto simp add: field_simps) hence "cor (1 - Z) = 2 / cor (1 + (cmod z)\<^sup>2)" by auto moreover have "X = 2 * Re(z) / (1 + (cmod z)\<^sup>2)" using X by (simp add: Re_stereographic) have "Y = 2 * Im(z) / (1 + (cmod z)\<^sup>2)" using Y by (simp add: Im_stereographic) have "X + \*Y = 2 * z / cor (1 + (cmod z)\<^sup>2)" using \1 + (cmod z)\<^sup>2 \ 0\ unfolding Complex_eq[of X Y, symmetric] by (subst \X = 2*Re(z) / (1 + (cmod z)\<^sup>2)\, subst \Y = 2*Im(z) / (1 + (cmod z)\<^sup>2)\, simp add: Complex_scale4 Complex_scale1) moreover have "1 + (cor (cmod (z1 / z2)))\<^sup>2 \ 0" by (rule one_plus_square_neq_zero) ultimately show ?thesis using \z2 \ 0\ \1 + (cmod z)\<^sup>2 \ 0\ by (simp, subst z)+ (rule_tac x="(2 / (1 + (cor (cmod (z1 / z2)))\<^sup>2)) / z2" in exI, auto) qed lemma stereographic_inv_stereographic [simp]: shows "stereographic (inv_stereographic w) = w" proof- have "w = stereographic (inv_stereographic w)" proof (transfer, transfer) fix w assume "w \ vec_zero" obtain w1 w2 where *: "w = (w1, w2)" by (cases w, auto) obtain x y z where **: "inv_stereographic_cvec_r3 w = (x, y, z)" by (cases "inv_stereographic_cvec_r3 w", blast) show "w \\<^sub>v stereographic_r3_cvec (inv_stereographic_cvec_r3 w)" using \w \ vec_zero\ stereographic_inv_stereographic'[of "w1/w2" w1 w2 x y z] * ** by (auto simp add: split_def Let_def split: if_split_asm) qed thus ?thesis by simp qed text \Stereographic projection is bijective function\ lemma bij_stereographic: shows "bij stereographic" unfolding bij_def inj_on_def surj_def proof (safe) fix a b assume "stereographic a = stereographic b" thus "a = b" proof (transfer, transfer) fix a b :: R3 obtain xa ya za xb yb zb where *: "a = (xa, ya, za)" "b = (xb, yb, zb)" by (cases a, cases b, auto) assume **: "a \ unit_sphere" "b \ unit_sphere" "stereographic_r3_cvec a \\<^sub>v stereographic_r3_cvec b" show "a = b" proof (cases "a = (0, 0, 1) \ b = (0, 0, 1)") case True thus ?thesis using * ** by (simp split: if_split_asm) force+ next case False then obtain k where ++: "k \ 0" "cor xb + \ * cor yb = k * (cor xa + \ * cor ya)" "1 - cor zb = k * (1 - cor za)" using * ** by (auto split: if_split_asm) { assume "xb + xa*zb = xa + xb*za" "yb + ya*zb = ya + yb*za" "xa*xa + ya*ya + za*za = 1" "xb*xb + yb*yb + zb*zb = 1" "za \ 1" "zb \ 1" hence "xa = xb \ ya = yb \ za = zb" by algebra } note *** = this have "za \ 1" "zb \ 1" using False * ** by auto have "k = (1 - cor zb) / (1 - cor za)" using \1 - cor zb = k * (1 - cor za)\ \za \ 1\ by simp hence "(1 - cor za) * (cor xb + \ * cor yb) = (1 - cor zb) * (cor xa + \ * cor ya)" using \za \ 1\ ++(2) by simp hence "xb + xa*zb = xa + xb*za" "yb + ya*zb = ya + yb*za" "xa*xa + ya*ya + za*za = 1" "xb*xb + yb*yb + zb*zb = 1" using * ** \za \ 1\ apply (simp_all add: field_simps) unfolding complex_of_real_def imaginary_unit.ctr by (simp_all add: legacy_Complex_simps) thus ?thesis using * ** *** \za \ 1\ \zb \ 1\ by simp qed qed next fix y show "\ x. y = stereographic x" by (rule_tac x="inv_stereographic y" in exI, simp) qed lemma inv_stereographic_stereographic [simp]: shows "inv_stereographic (stereographic x) = x" using stereographic_inv_stereographic[of "stereographic x"] using bij_stereographic unfolding bij_def inj_on_def by simp lemma inv_stereographic_is_inv: shows "inv_stereographic = inv stereographic" by (rule inv_equality[symmetric], simp_all) (* ----------------------------------------------------------------- *) subsection \Circles on the sphere\ (* ----------------------------------------------------------------- *) text \Circlines in the plane correspond to circles on the Riemann sphere, and we formally establish this connection. Every circle in three--dimensional space can be obtained as the intersection of a sphere and a plane. We establish a one-to-one correspondence between circles on the Riemann sphere and planes in space. Note that the plane need not intersect the sphere, but we will still say that it defines a single imaginary circle. However, for one special circline (the one with the identity representative matrix), there does not exist a plane in $\mathbb{R}^3$ that would correspond to it --- in order to have this, instead of considering planes in $\mathbb{R}^3$, we must consider three dimensional projective space and consider the infinite (hyper)plane.\ text \Planes in $R^3$ are given by equations $ax+by+cz=d$. Two four-tuples of coefficients $(a, b, c, d)$ give the same plane iff they are proportional.\ type_synonym R4 = "real \ real \ real \ real" fun mult_sv :: "real \ R4 \ R4" (infixl "*\<^sub>s\<^sub>v\<^sub>4" 100) where "k *\<^sub>s\<^sub>v\<^sub>4 (a, b, c, d) = (k*a, k*b, k*c, k*d)" abbreviation plane_vectors where "plane_vectors \ {(a::real, b::real, c::real, d::real). a \ 0 \ b \ 0 \ c \ 0 \ d \ 0}" typedef plane_vec = "plane_vectors" by (rule_tac x="(1, 1, 1, 1)" in exI) simp setup_lifting type_definition_plane_vec definition plane_vec_eq_r4 :: "R4 \ R4 \ bool" where [simp]: "plane_vec_eq_r4 v1 v2 \ (\ k. k \ 0 \ v2 = k *\<^sub>s\<^sub>v\<^sub>4 v1)" lift_definition plane_vec_eq :: "plane_vec \ plane_vec \ bool" is plane_vec_eq_r4 done lemma mult_sv_one [simp]: shows "1 *\<^sub>s\<^sub>v\<^sub>4 x = x" by (cases x) simp lemma mult_sv_distb [simp]: shows "x *\<^sub>s\<^sub>v\<^sub>4 (y *\<^sub>s\<^sub>v\<^sub>4 v) = (x*y) *\<^sub>s\<^sub>v\<^sub>4 v" by (cases v) simp quotient_type plane = plane_vec / plane_vec_eq proof (rule equivpI) show "reflp plane_vec_eq" unfolding reflp_def by (auto simp add: plane_vec_eq_def) (rule_tac x="1" in exI, simp) next show "symp plane_vec_eq" unfolding symp_def by (auto simp add: plane_vec_eq_def) (rule_tac x="1/k" in exI, simp) next show "transp plane_vec_eq" unfolding transp_def by (auto simp add: plane_vec_eq_def) (rule_tac x="ka*k" in exI, simp) qed text \Plane coefficients give a linear equation and the point on the Riemann sphere lies on the circle determined by the plane iff its representation satisfies that linear equation.\ definition on_sphere_circle_r4_r3 :: "R4 \ R3 \ bool" where [simp]: "on_sphere_circle_r4_r3 \ A \ (let (X, Y, Z) = A; (a, b, c, d) = \ in a*X + b*Y + c*Z + d = 0)" lift_definition on_sphere_circle_vec :: "plane_vec \ R3 \ bool" is on_sphere_circle_r4_r3 done lift_definition on_sphere_circle :: "plane \ riemann_sphere \ bool" is on_sphere_circle_vec proof (transfer) fix pv1 pv2 :: R4 and w :: R3 obtain a1 b1 c1 d1 a2 b2 c2 d2 x y z where *: "pv1 = (a1, b1, c1, d1)" "pv2 = (a2, b2, c2, d2)" "w = (x, y, z)" by (cases pv1, cases pv2, cases w, auto) assume "pv1 \ plane_vectors" "pv2 \ plane_vectors" "w \ unit_sphere" "plane_vec_eq_r4 pv1 pv2" then obtain k where **: "a2 = k*a1" "b2 = k*b1" "c2 = k*c1" "d2 = k*d1" "k \ 0" using * by auto have "k * a1 * x + k * b1 * y + k * c1 * z + k * d1 = k*(a1*x + b1*y + c1*z + d1)" by (simp add: field_simps) thus "on_sphere_circle_r4_r3 pv1 w = on_sphere_circle_r4_r3 pv2 w" using * ** by simp qed definition sphere_circle_set where "sphere_circle_set \ = {A. on_sphere_circle \ A}" (* ----------------------------------------------------------------- *) subsection \Connections of circlines in the plane and circles on the Riemann sphere\ (* ----------------------------------------------------------------- *) text \We introduce stereographic and inverse stereographic projection between circles on the Riemann sphere and circlines in the extended complex plane.\ definition inv_stereographic_circline_cmat_r4 :: "complex_mat \ R4" where [simp]: "inv_stereographic_circline_cmat_r4 H = (let (A, B, C, D) = H in (Re (B+C), Re(\*(C-B)), Re(A-D), Re(D+A)))" lift_definition inv_stereographic_circline_clmat_pv :: "circline_mat \ plane_vec" is inv_stereographic_circline_cmat_r4 by (auto simp add: hermitean_def mat_adj_def mat_cnj_def real_imag_0 eq_cnj_iff_real) lift_definition inv_stereographic_circline :: "circline \ plane" is inv_stereographic_circline_clmat_pv apply transfer apply simp apply (erule exE) apply (rule_tac x="k" in exI) apply (case_tac "circline_mat1", case_tac "circline_mat2") apply (simp add: field_simps) done definition stereographic_circline_r4_cmat :: "R4 \ complex_mat" where [simp]: "stereographic_circline_r4_cmat \ = (let (a, b, c, d) = \ in (cor ((c+d)/2) , ((cor a + \ * cor b)/2), ((cor a - \ * cor b)/2), cor ((d-c)/2)))" lift_definition stereographic_circline_pv_clmat :: "plane_vec \ circline_mat" is stereographic_circline_r4_cmat by (auto simp add: hermitean_def mat_adj_def mat_cnj_def) lift_definition stereographic_circline :: "plane \ circline" is stereographic_circline_pv_clmat apply transfer apply transfer apply (case_tac plane_vec1, case_tac plane_vec2, simp, erule exE, rule_tac x=k in exI, simp add: field_simps) done text \Stereographic and inverse stereographic projection of circlines are mutually inverse.\ lemma stereographic_circline_inv_stereographic_circline: shows "stereographic_circline \ inv_stereographic_circline = id" proof (rule ext, simp) fix H show "stereographic_circline (inv_stereographic_circline H) = H" proof (transfer, transfer) fix H assume hh: "hermitean H \ H \ mat_zero" obtain A B C D where HH: "H = (A, B, C, D)" by (cases "H") auto have "is_real A" "is_real D" "C = cnj B" using HH hh hermitean_elems[of A B C D] by auto thus "circline_eq_cmat (stereographic_circline_r4_cmat (inv_stereographic_circline_cmat_r4 H)) H" using HH apply simp apply (rule_tac x=1 in exI, cases B) by (smt add_uminus_conv_diff complex_cnj_add complex_cnj_complex_of_real complex_cnj_i complex_cnj_mult complex_cnj_one complex_eq distrib_left_numeral mult.commute mult.left_commute mult.left_neutral mult_cancel_right2 mult_minus_left of_real_1 one_add_one) qed qed text \Stereographic and inverse stereographic projection of circlines are mutually inverse.\ lemma inv_stereographic_circline_stereographic_circline: "inv_stereographic_circline \ stereographic_circline = id" proof (rule ext, simp) fix \ show "inv_stereographic_circline (stereographic_circline \) = \" proof (transfer, transfer) fix \ assume aa: "\ \ plane_vectors" obtain a b c d where AA: "\ = (a, b, c, d)" by (cases "\") auto thus "plane_vec_eq_r4 (inv_stereographic_circline_cmat_r4 (stereographic_circline_r4_cmat \)) \" using AA by simp (rule_tac x=1 in exI, auto simp add: field_simps complex_of_real_def) qed qed lemma stereographic_sphere_circle_set'': shows "on_sphere_circle (inv_stereographic_circline H) z \ on_circline H (stereographic z)" proof (transfer, transfer) fix M :: R3 and H :: complex_mat assume hh: "hermitean H \ H \ mat_zero" "M \ unit_sphere" obtain A B C D where HH: "H = (A, B, C, D)" by (cases "H") auto have *: "is_real A" "is_real D" "C = cnj B" using hh HH hermitean_elems[of A B C D] by auto obtain x y z where MM: "M = (x, y, z)" by (cases "M") auto show "on_sphere_circle_r4_r3 (inv_stereographic_circline_cmat_r4 H) M \ on_circline_cmat_cvec H (stereographic_r3_cvec M)" (is "?lhs = ?rhs") proof assume ?lhs show ?rhs proof (cases "z=1") case True hence "x = 0" "y = 0" using MM hh by auto thus ?thesis using * \?lhs\ HH MM \z=1\ by (cases A, simp add: vec_cnj_def Complex_eq Let_def) next case False hence "Re A*(1+z) + 2*Re B*x + 2*Im B*y + Re D*(1-z) = 0" using * \?lhs\ HH MM by (simp add: Let_def field_simps) hence "(Re A*(1+z) + 2*Re B*x + 2*Im B*y + Re D*(1-z))*(1-z) = 0" by simp hence "Re A*(1+z)*(1-z) + 2*Re B*x*(1-z) + 2*Im B*y*(1-z) + Re D*(1-z)*(1-z) = 0" by (simp add: field_simps) moreover have "x*x+y*y = (1+z)*(1-z)" using MM hh by (simp add: field_simps) ultimately have "Re A*(x*x+y*y) + 2*Re B*x*(1-z) + 2*Im B*y*(1-z) + Re D*(1-z)*(1-z) = 0" by simp hence "(x * Re A + (1 - z) * Re B) * x - (- (y * Re A) + - ((1 - z) * Im B)) * y + (x * Re B + y * Im B + (1 - z) * Re D) * (1 - z) = 0" by (simp add: field_simps) thus ?thesis using \z \ 1\ HH MM * \Re A*(1+z) + 2*Re B*x + 2*Im B*y + Re D*(1-z) = 0\ apply (simp add: Let_def vec_cnj_def) apply (subst complex_eq_iff) apply (simp add: field_simps) done qed next assume ?rhs show ?lhs proof (cases "z=1") case True hence "x = 0" "y = 0" using MM hh by auto thus ?thesis using HH MM \?rhs\ \z = 1\ by (simp add: Let_def vec_cnj_def) next case False hence "(x * Re A + (1 - z) * Re B) * x - (- (y * Re A) + - ((1 - z) * Im B)) * y + (x * Re B + y * Im B + (1 - z) * Re D) * (1 - z) = 0" using HH MM * \?rhs\ by (simp add: Let_def vec_cnj_def complex_eq_iff) hence "Re A*(x*x+y*y) + 2*Re B*x*(1-z) + 2*Im B*y*(1-z) + Re D*(1-z)*(1-z) = 0" by (simp add: field_simps) moreover have "x*x + y*y = (1+z)*(1-z)" using MM hh by (simp add: field_simps) ultimately have "Re A*(1+z)*(1-z) + 2*Re B*x*(1-z) + 2*Im B*y*(1-z) + Re D*(1-z)*(1-z) = 0" by simp hence "(Re A*(1+z) + 2*Re B*x + 2*Im B*y + Re D*(1-z))*(1-z) = 0" by (simp add: field_simps) hence "Re A*(1+z) + 2*Re B*x + 2*Im B*y + Re D*(1-z) = 0" using \z \ 1\ by simp thus ?thesis using MM HH * by (simp add: field_simps) qed qed qed lemma stereographic_sphere_circle_set' [simp]: shows "stereographic ` sphere_circle_set (inv_stereographic_circline H) = circline_set H" unfolding sphere_circle_set_def circline_set_def apply safe proof- fix x assume "on_sphere_circle (inv_stereographic_circline H) x" thus "on_circline H (stereographic x)" using stereographic_sphere_circle_set'' by simp next fix x assume "on_circline H x" show "x \ stereographic ` {z. on_sphere_circle (inv_stereographic_circline H) z}" proof show "x = stereographic (inv_stereographic x)" by simp next show "inv_stereographic x \ {z. on_sphere_circle (inv_stereographic_circline H) z}" using stereographic_sphere_circle_set''[of H "inv_stereographic x"] \on_circline H x\ by simp qed qed text \The projection of the set of points on a circle on the Riemann sphere is exactly the set of points on the circline obtained by the just introduced circle stereographic projection.\ lemma stereographic_sphere_circle_set: shows "stereographic ` sphere_circle_set H = circline_set (stereographic_circline H)" using stereographic_sphere_circle_set'[of "stereographic_circline H"] using inv_stereographic_circline_stereographic_circline unfolding comp_def by (metis id_apply) text \Stereographic projection of circlines is bijective.\ lemma bij_stereographic_circline: shows "bij stereographic_circline" using stereographic_circline_inv_stereographic_circline inv_stereographic_circline_stereographic_circline using o_bij by blast text \Inverse stereographic projection is bijective.\ lemma bij_inv_stereographic_circline: shows "bij inv_stereographic_circline" using stereographic_circline_inv_stereographic_circline inv_stereographic_circline_stereographic_circline using o_bij by blast end diff --git a/thys/Complex_Geometry/Unitary11_Matrices.thy b/thys/Complex_Geometry/Unitary11_Matrices.thy --- a/thys/Complex_Geometry/Unitary11_Matrices.thy +++ b/thys/Complex_Geometry/Unitary11_Matrices.thy @@ -1,607 +1,607 @@ (* ----------------------------------------------------------------- *) subsection \Generalized unitary matrices with signature $(1, 1)$\ (* ----------------------------------------------------------------- *) theory Unitary11_Matrices imports Matrices More_Complex begin text \ When acting as Möbius transformations in the extended complex plane, generalized complex $2\times 2$ unitary matrices fix the imaginary unit circle (a Hermitean form with (2, 0) signature). We now describe matrices that fix the ordinary unit circle (a Hermitean form with (1, 1) signature, i.e., one positive and one negative element on the diagonal). These are extremely important for further formalization, since they will represent disc automorphisims and isometries of the Poincar\'e disc. The development of this theory follows the development of the theory of generalized unitary matrices. \ text \Unitary11 matrices\ definition unitary11 where "unitary11 M \ congruence M (1, 0, 0, -1) = (1, 0, 0, -1)" text \Generalized unitary11 matrices\ definition unitary11_gen where "unitary11_gen M \ (\ k. k \ 0 \ congruence M (1, 0, 0, -1) = k *\<^sub>s\<^sub>m (1, 0, 0, -1))" text \Scalar can always be a non-zero real number\ lemma unitary11_gen_real: shows "unitary11_gen M \ (\ k. k \ 0 \ congruence M (1, 0, 0, -1) = cor k *\<^sub>s\<^sub>m (1, 0, 0, -1))" unfolding unitary11_gen_def proof (auto simp del: congruence_def) fix k assume "k \ 0" "congruence M (1, 0, 0, -1) = (k, 0, 0, - k)" hence "mat_det (congruence M (1, 0, 0, -1)) = -k*k" by simp moreover have "is_real (mat_det (congruence M (1, 0, 0, -1)))" "Re (mat_det (congruence M (1, 0, 0, -1))) \ 0" by (auto simp add: mat_det_adj) ultimately have "is_real (k*k)" "Re (-k*k) \ 0" by auto hence "is_real (k*k) \ Re (k * k) > 0" using \k \ 0\ by (smt complex_eq_if_Re_eq mult_eq_0_iff mult_minus_left uminus_complex.simps(1) zero_complex.simps(1) zero_complex.simps(2)) hence "is_real k" by auto thus "\ka. ka \ 0 \ k = cor ka" using \k \ 0\ by (rule_tac x="Re k" in exI) (cases k, auto simp add: Complex_eq) qed text \Unitary11 matrices are special cases of generalized unitary 11 matrices\ lemma unitary11_unitary11_gen [simp]: assumes "unitary11 M" shows "unitary11_gen M" using assms unfolding unitary11_gen_def unitary11_def by (rule_tac x="1" in exI, auto) text \All generalized unitary11 matrices are regular\ lemma unitary11_gen_regular: assumes "unitary11_gen M" shows "mat_det M \ 0" proof- from assms obtain k where "k \ 0" "mat_adj M *\<^sub>m\<^sub>m (1, 0, 0, -1) *\<^sub>m\<^sub>m M = cor k *\<^sub>s\<^sub>m (1, 0, 0, -1)" unfolding unitary11_gen_real by auto hence "mat_det (mat_adj M *\<^sub>m\<^sub>m (1, 0, 0, -1) *\<^sub>m\<^sub>m M) \ 0" by simp thus ?thesis by (simp add: mat_det_adj) qed lemmas unitary11_regular = unitary11_gen_regular[OF unitary11_unitary11_gen] (* ----------------------------------------------------------------- *) subsubsection \The characterization in terms of matrix elements\ (* ----------------------------------------------------------------- *) text \Special matrices are those having the determinant equal to 1. We first give their characterization.\ lemma unitary11_special: assumes "unitary11 M" and "mat_det M = 1" shows "\ a b. M = (a, b, cnj b, cnj a)" proof- have "mat_adj M *\<^sub>m\<^sub>m (1, 0, 0, -1) = (1, 0, 0, -1) *\<^sub>m\<^sub>m mat_inv M" using assms mult_mm_inv_r by (simp add: unitary11_def) thus ?thesis using assms(2) by (cases M) (simp add: mat_adj_def mat_cnj_def) qed lemma unitary11_gen_special: assumes "unitary11_gen M" and "mat_det M = 1" shows "\ a b. M = (a, b, cnj b, cnj a) \ M = (a, b, -cnj b, -cnj a)" proof- from assms obtain k where *: "k \ 0" "mat_adj M *\<^sub>m\<^sub>m (1, 0, 0, -1) *\<^sub>m\<^sub>m M = cor k *\<^sub>s\<^sub>m (1, 0, 0, -1)" unfolding unitary11_gen_real by auto hence "mat_det (mat_adj M *\<^sub>m\<^sub>m (1, 0, 0, -1) *\<^sub>m\<^sub>m M) = - cor k* cor k" by simp hence "mat_det (mat_adj M *\<^sub>m\<^sub>m M) = cor k* cor k" by simp hence "cor k* cor k = 1" using assms(2) by (simp add: mat_det_adj) hence "cor k = 1 \ cor k = -1" using square_eq_1_iff[of "cor k"] by simp moreover have "mat_adj M *\<^sub>m\<^sub>m (1, 0, 0, -1) = (cor k *\<^sub>s\<^sub>m (1, 0, 0, -1)) *\<^sub>m\<^sub>m mat_inv M " using * using assms mult_mm_inv_r mat_eye_r mat_eye_l by auto moreover obtain a b c d where "M = (a, b, c, d)" by (cases M) auto ultimately have "M = (a, b, cnj b, cnj a) \ M = (a, b, -cnj b, -cnj a)" using assms(2) by (auto simp add: mat_adj_def mat_cnj_def) thus ?thesis by auto qed text \A characterization of all generalized unitary11 matrices\ lemma unitary11_gen_iff': shows "unitary11_gen M \ (\ a b k. k \ 0 \ mat_det (a, b, cnj b, cnj a) \ 0 \ (M = k *\<^sub>s\<^sub>m (a, b, cnj b, cnj a) \ M = k *\<^sub>s\<^sub>m (-1, 0, 0, 1) *\<^sub>m\<^sub>m (a, b, cnj b, cnj a)))" (is "?lhs = ?rhs") proof assume ?lhs obtain d where *: "d*d = mat_det M" using ex_complex_sqrt by auto hence "d \ 0" using unitary11_gen_regular[OF \unitary11_gen M\] by auto from \unitary11_gen M\ obtain k where "k \ 0" "mat_adj M *\<^sub>m\<^sub>m (1, 0, 0, -1) *\<^sub>m\<^sub>m M = cor k *\<^sub>s\<^sub>m (1, 0, 0, -1)" unfolding unitary11_gen_real by auto hence "mat_adj ((1/d)*\<^sub>s\<^sub>mM)*\<^sub>m\<^sub>m (1, 0, 0, -1) *\<^sub>m\<^sub>m ((1/d)*\<^sub>s\<^sub>mM) = (cor k / (d*cnj d)) *\<^sub>s\<^sub>m (1, 0, 0, -1)" by simp moreover have "is_real (cor k / (d * cnj d))" by (metis complex_In_mult_cnj_zero div_reals Im_complex_of_real) hence "cor (Re (cor k / (d * cnj d))) = cor k / (d * cnj d)" by simp ultimately have "unitary11_gen ((1/d)*\<^sub>s\<^sub>mM)" unfolding unitary11_gen_real using \d \ 0\ \k \ 0\ using \cor (Re (cor k / (d * cnj d))) = cor k / (d * cnj d)\ by (rule_tac x="Re (cor k / (d * cnj d))" in exI, auto, simp add: *) moreover have "mat_det ((1 / d) *\<^sub>s\<^sub>m M) = 1" using * unitary11_gen_regular[of M] \unitary11_gen M\ by auto ultimately obtain a b where "(a, b, cnj b, cnj a) = (1 / d) *\<^sub>s\<^sub>m M \ (a, b, -cnj b, -cnj a) = (1 / d) *\<^sub>s\<^sub>m M" using unitary11_gen_special[of "(1 / d) *\<^sub>s\<^sub>m M"] by force thus ?rhs proof assume "(a, b, cnj b, cnj a) = (1 / d) *\<^sub>s\<^sub>m M" moreover hence "mat_det (a, b, cnj b, cnj a) \ 0" using unitary11_gen_regular[OF \unitary11_gen M\] \d \ 0\ by auto ultimately show ?rhs using \d \ 0\ by (rule_tac x="a" in exI, rule_tac x="b" in exI, rule_tac x="d" in exI, simp) next assume *: "(a, b, -cnj b, -cnj a) = (1 / d) *\<^sub>s\<^sub>m M" hence " (1 / d) *\<^sub>s\<^sub>m M = (a, b, -cnj b, -cnj a)" by simp hence "M = (a * d, b * d, - (d * cnj b), - (d * cnj a))" using \d \ 0\ using mult_sm_inv_l[of "1/d" M "(a, b, -cnj b, -cnj a)", symmetric] by (simp add: field_simps) moreover have "mat_det (a, b, -cnj b, -cnj a) \ 0" using * unitary11_gen_regular[OF \unitary11_gen M\] \d \ 0\ by auto ultimately show ?thesis using \d \ 0\ by (rule_tac x="a" in exI, rule_tac x="b" in exI, rule_tac x="-d" in exI) (simp add: field_simps) qed next assume ?rhs then obtain a b k where "k \ 0" "mat_det (a, b, cnj b, cnj a) \ 0" "M = k *\<^sub>s\<^sub>m (a, b, cnj b, cnj a) \ M = k *\<^sub>s\<^sub>m (-1, 0, 0, 1) *\<^sub>m\<^sub>m (a, b, cnj b, cnj a)" by auto moreover let ?x = "cnj k * cnj a * (k * a) + - (cnj k * b * (k * cnj b))" have "?x = (k*cnj k)*(a*cnj a - b*cnj b)" by (auto simp add: field_simps) hence "is_real ?x" by simp hence "cor (Re ?x) = ?x" by (rule complex_of_real_Re) moreover have "?x \ 0" using mult_eq_0_iff[of "cnj k * k" "(cnj a * a + - cnj b * b)"] using \mat_det (a, b, cnj b, cnj a) \ 0\ \k \ 0\ by (auto simp add: field_simps) hence "Re ?x \ 0" using \is_real ?x\ by (metis calculation(4) of_real_0) ultimately show ?lhs unfolding unitary11_gen_real by (rule_tac x="Re ?x" in exI) (auto simp add: mat_adj_def mat_cnj_def) qed text \Another characterization of all generalized unitary11 matrices. They are products of rotation and Blaschke factor matrices.\ lemma unitary11_gen_cis_blaschke: assumes "k \ 0" and "M = k *\<^sub>s\<^sub>m (a, b, cnj b, cnj a)" and "a \ 0" and "mat_det (a, b, cnj b, cnj a) \ 0" shows "\ k' \ a'. k' \ 0 \ a' * cnj a' \ 1 \ M = k' *\<^sub>s\<^sub>m (cis \, 0, 0, 1) *\<^sub>m\<^sub>m (1, -a', -cnj a', 1)" proof- have "a = cnj a * cis (2 * arg a)" using rcis_cmod_arg[of a] rcis_cnj[of a] using cis_rcis_eq rcis_mult by simp thus ?thesis using assms by (rule_tac x="k*cnj a" in exI, rule_tac x="2*arg a" in exI, rule_tac x="- b / a" in exI) (auto simp add: field_simps) qed lemma unitary11_gen_cis_blaschke': assumes "k \ 0" and "M = k *\<^sub>s\<^sub>m (-1, 0, 0, 1) *\<^sub>m\<^sub>m (a, b, cnj b, cnj a)" and "a \ 0" and "mat_det (a, b, cnj b, cnj a) \ 0" shows "\ k' \ a'. k' \ 0 \ a' * cnj a' \ 1 \ M = k' *\<^sub>s\<^sub>m (cis \, 0, 0, 1) *\<^sub>m\<^sub>m (1, -a', -cnj a', 1)" proof- obtain k' \ a' where *: "k' \ 0" "k *\<^sub>s\<^sub>m (a, b, cnj b, cnj a) = k' *\<^sub>s\<^sub>m (cis \, 0, 0, 1) *\<^sub>m\<^sub>m (1, -a', -cnj a', 1)" "a' * cnj a' \ 1" using unitary11_gen_cis_blaschke[OF \k \ 0\ _ \a \ 0\] \mat_det (a, b, cnj b, cnj a) \ 0\ by blast have "(cis \, 0, 0, 1) *\<^sub>m\<^sub>m (-1, 0, 0, 1) = (cis (\ + pi), 0, 0, 1)" by (simp add: cis_def complex.corec Complex_eq) thus ?thesis using * \M = k *\<^sub>s\<^sub>m (-1, 0, 0, 1) *\<^sub>m\<^sub>m (a, b, cnj b, cnj a)\ by (rule_tac x="k'" in exI, rule_tac x="\ + pi" in exI, rule_tac x="a'" in exI, simp) qed lemma unitary11_gen_cis_blaschke_rev: assumes "k' \ 0" and "M = k' *\<^sub>s\<^sub>m (cis \, 0, 0, 1) *\<^sub>m\<^sub>m (1, -a', -cnj a', 1)" and "a' * cnj a' \ 1" shows "\ k a b. k \ 0 \ mat_det (a, b, cnj b, cnj a) \ 0 \ M = k *\<^sub>s\<^sub>m (a, b, cnj b, cnj a)" using assms apply (rule_tac x="k'*cis(\/2)" in exI, rule_tac x="cis(\/2)" in exI, rule_tac x="-a'*cis(\/2)" in exI) apply (simp add: cis_mult mult.commute mult.left_commute) done lemma unitary11_gen_cis_inversion: assumes "k \ 0" and "M = k *\<^sub>s\<^sub>m (0, b, cnj b, 0)" and "b \ 0" shows "\ k' \. k' \ 0 \ M = k' *\<^sub>s\<^sub>m (cis \, 0, 0, 1) *\<^sub>m\<^sub>m (0, 1, 1, 0)" using assms using rcis_cmod_arg[of b, symmetric] rcis_cnj[of b] cis_rcis_eq by simp (rule_tac x="2*arg b" in exI, simp add: rcis_mult) lemma unitary11_gen_cis_inversion': assumes "k \ 0" and "M = k *\<^sub>s\<^sub>m (-1, 0, 0, 1) *\<^sub>m\<^sub>m (0, b, cnj b, 0)" and "b \ 0" shows "\ k' \. k' \ 0 \ M = k' *\<^sub>s\<^sub>m (cis \, 0, 0, 1) *\<^sub>m\<^sub>m (0, 1, 1, 0)" proof- obtain k' \ where *: "k' \ 0" "k *\<^sub>s\<^sub>m (0, b, cnj b, 0) = k' *\<^sub>s\<^sub>m (cis \, 0, 0, 1) *\<^sub>m\<^sub>m (0, 1, 1, 0)" using unitary11_gen_cis_inversion[OF \k \ 0\ _ \b \ 0\] by metis have "(cis \, 0, 0, 1) *\<^sub>m\<^sub>m (-1, 0, 0, 1) = (cis (\ + pi), 0, 0, 1)" by (simp add: cis_def complex.corec Complex_eq) thus ?thesis using * \M = k *\<^sub>s\<^sub>m (-1, 0, 0, 1) *\<^sub>m\<^sub>m (0, b, cnj b, 0)\ by (rule_tac x="k'" in exI, rule_tac x="\ + pi" in exI, simp) qed lemma unitary11_gen_cis_inversion_rev: assumes "k' \ 0" and "M = k' *\<^sub>s\<^sub>m (cis \, 0, 0, 1) *\<^sub>m\<^sub>m (0, 1, 1, 0)" shows "\ k a b. k \ 0 \ mat_det (a, b, cnj b, cnj a) \ 0 \ M = k *\<^sub>s\<^sub>m (a, b, cnj b, cnj a)" using assms by (rule_tac x="k'*cis(\/2)" in exI, rule_tac x=0 in exI, rule_tac x="cis(\/2)" in exI) (simp add: cis_mult) text \Another characterization of generalized unitary11 matrices\ lemma unitary11_gen_iff: shows "unitary11_gen M \ (\ k a b. k \ 0 \ mat_det (a, b, cnj b, cnj a) \ 0 \ M = k *\<^sub>s\<^sub>m (a, b, cnj b, cnj a))" (is "?lhs = ?rhs") proof assume ?lhs then obtain a b k where *: "k \ 0" "mat_det (a, b, cnj b, cnj a) \ 0" "M = k *\<^sub>s\<^sub>m (a, b, cnj b, cnj a) \ M = k *\<^sub>s\<^sub>m (-1, 0, 0, 1) *\<^sub>m\<^sub>m (a, b, cnj b, cnj a)" using unitary11_gen_iff' by auto show ?rhs proof (cases "M = k *\<^sub>s\<^sub>m (a, b, cnj b, cnj a)") case True thus ?thesis using * by auto next case False hence **: "M = k *\<^sub>s\<^sub>m (-1, 0, 0, 1) *\<^sub>m\<^sub>m (a, b, cnj b, cnj a)" using * by simp show ?thesis proof (cases "a = 0") case True hence "b \ 0" using * by auto show ?thesis using unitary11_gen_cis_inversion_rev[of _ M] using ** \a = 0\ using unitary11_gen_cis_inversion'[OF \k \ 0\ _ \b \ 0\, of M] by auto next case False show ?thesis using unitary11_gen_cis_blaschke_rev[of _ M] using ** using unitary11_gen_cis_blaschke'[OF \k \ 0\ _ \a \ 0\, of M b] \mat_det (a, b, cnj b, cnj a) \ 0\ by blast qed qed next assume ?rhs thus ?lhs using unitary11_gen_iff' by auto qed lemma unitary11_iff: shows "unitary11 M \ (\ a b k. (cmod a)\<^sup>2 > (cmod b)\<^sup>2 \ (cmod k)\<^sup>2 = 1 / ((cmod a)\<^sup>2 - (cmod b)\<^sup>2) \ M = k *\<^sub>s\<^sub>m (a, b, cnj b, cnj a))" (is "?lhs = ?rhs") proof assume ?lhs obtain k a b where *: "M = k *\<^sub>s\<^sub>m (a, b, cnj b, cnj a)""mat_det (a, b, cnj b, cnj a) \ 0" "k \ 0" using unitary11_gen_iff unitary11_unitary11_gen[OF \unitary11 M\] by auto have md: "mat_det (a, b, cnj b, cnj a) = cor ((cmod a)\<^sup>2 - (cmod b)\<^sup>2)" by (auto simp add: complex_mult_cnj_cmod) hence **: "(cmod a)\<^sup>2 \ (cmod b)\<^sup>2" using \mat_det (a, b, cnj b, cnj a) \ 0\ by auto have "k * cnj k * mat_det (a, b, cnj b, cnj a) = 1" using \M = k *\<^sub>s\<^sub>m (a, b, cnj b, cnj a)\ using \unitary11 M\ unfolding unitary11_def by (auto simp add: mat_adj_def mat_cnj_def) (simp add: field_simps) hence ***: "(cmod k)\<^sup>2 * ((cmod a)\<^sup>2 - (cmod b)\<^sup>2) = 1" - by (subst (asm) complex_mult_cnj_cmod, subst (asm) md, subst (asm) cor_mult[symmetric]) (metis of_real_1 of_real_eq_iff) + by (metis complex_mult_cnj_cmod md of_real_1 of_real_eq_iff of_real_mult) hence "((cmod a)\<^sup>2 - (cmod b)\<^sup>2) = 1 / (cmod k)\<^sup>2" by (cases "k=0") (auto simp add: field_simps) hence "cmod a ^ 2 = cmod b ^ 2 + 1 / cmod k ^ 2" by simp thus ?rhs using \M = k *\<^sub>s\<^sub>m (a, b, cnj b, cnj a)\ ** mat_eye_l by (rule_tac x="a" in exI, rule_tac x="b" in exI, rule_tac x="k" in exI) (auto simp add: complex_mult_cnj_cmod intro!: ) next assume ?rhs then obtain a b k where "(cmod b)\<^sup>2 < (cmod a)\<^sup>2 \ (cmod k)\<^sup>2 = 1 / ((cmod a)\<^sup>2 - (cmod b)\<^sup>2) \ M = k *\<^sub>s\<^sub>m (a, b, cnj b, cnj a)" by auto moreover have "cnj k * cnj a * (k * a) + - (cnj k * b * (k * cnj b)) = (cor ((cmod k)\<^sup>2 * ((cmod a)\<^sup>2 - (cmod b)\<^sup>2)))" proof- have "cnj k * cnj a * (k * a) = cor ((cmod k)\<^sup>2 * (cmod a)\<^sup>2)" using complex_mult_cnj_cmod[of a] complex_mult_cnj_cmod[of k] by (auto simp add: field_simps) moreover have "cnj k * b * (k * cnj b) = cor ((cmod k)\<^sup>2 * (cmod b)\<^sup>2)" using complex_mult_cnj_cmod[of b, symmetric] complex_mult_cnj_cmod[of k] by (auto simp add: field_simps) ultimately show ?thesis by (auto simp add: field_simps) qed ultimately show ?lhs unfolding unitary11_def by (auto simp add: mat_adj_def mat_cnj_def field_simps) qed (* ----------------------------------------------------------------- *) subsubsection \Group properties\ (* ----------------------------------------------------------------- *) text \Generalized unitary11 matrices form a group under multiplication (it is sometimes denoted by $GU_{1, 1}(2, \mathbb{C})$). The group is also closed under non-zero complex scalar multiplication. Since these matrices are always regular, they form a subgroup of general linear group (usually denoted by $GL(2, \mathbb{C})$) of all regular matrices.\ lemma unitary11_gen_mult_sm: assumes "k \ 0" and "unitary11_gen M" shows "unitary11_gen (k *\<^sub>s\<^sub>m M)" proof- have "k * cnj k = cor (Re (k * cnj k))" by (subst complex_of_real_Re) auto thus ?thesis using assms unfolding unitary11_gen_real by auto (rule_tac x="Re (k*cnj k) * ka" in exI, auto) qed lemma unitary11_gen_div_sm: assumes "k \ 0" and "unitary11_gen (k *\<^sub>s\<^sub>m M)" shows "unitary11_gen M" using assms unitary11_gen_mult_sm[of "1/k" "k *\<^sub>s\<^sub>m M"] by simp lemma unitary11_inv: assumes "k \ 0" and "M = k *\<^sub>s\<^sub>m (a, b, cnj b, cnj a)" and "mat_det (a, b, cnj b, cnj a) \ 0" shows "\ k' a' b'. k' \ 0 \ mat_inv M = k' *\<^sub>s\<^sub>m (a', b', cnj b', cnj a') \ mat_det (a', b', cnj b', cnj a') \ 0" using assms by (subst assms, subst mat_inv_mult_sm[OF assms(1)]) (rule_tac x="1/(k * mat_det (a, b, cnj b, cnj a))" in exI, rule_tac x="cnj a" in exI, rule_tac x="-b" in exI, simp add: field_simps) lemma unitary11_comp: assumes "k1 \ 0" and "M1 = k1 *\<^sub>s\<^sub>m (a1, b1, cnj b1, cnj a1)" and "mat_det (a1, b1, cnj b1, cnj a1) \ 0" "k2 \ 0" "M2 = k2 *\<^sub>s\<^sub>m (a2, b2, cnj b2, cnj a2)" "mat_det (a2, b2, cnj b2, cnj a2) \ 0" shows "\ k a b. k \ 0 \ M1 *\<^sub>m\<^sub>m M2 = k *\<^sub>s\<^sub>m (a, b, cnj b, cnj a) \ mat_det (a, b, cnj b, cnj a) \ 0" using assms apply (rule_tac x="k1*k2" in exI) apply (rule_tac x="a1*a2 + b1*cnj b2" in exI) apply (rule_tac x="a1*b2 + b1*cnj a2" in exI) proof (auto simp add: algebra_simps) assume *: "a1 * (a2 * (cnj a1 * cnj a2)) + b1 * (b2 * (cnj b1 * cnj b2)) = a1 * (b2 * (cnj a1 * cnj b2)) + a2 * (b1 * (cnj a2 * cnj b1))" and **: "a1*cnj a1 \ b1 * cnj b1" "a2*cnj a2 \ b2*cnj b2" hence "(a1*cnj a1)*(a2*cnj a2 - b2*cnj b2) = (b1*cnj b1)*(a2*cnj a2 - b2*cnj b2)" by (simp add: field_simps) hence "a1*cnj a1 = b1*cnj b1" using **(2) by simp thus False using **(1) by simp qed lemma unitary11_gen_mat_inv: assumes "unitary11_gen M" and "mat_det M \ 0" shows "unitary11_gen (mat_inv M)" proof- obtain k a b where "k \ 0 \ mat_det (a, b, cnj b, cnj a) \ 0 \ M = k *\<^sub>s\<^sub>m (a, b, cnj b, cnj a)" using assms unitary11_gen_iff[of M] by auto then obtain k' a' b' where "k' \ 0 \ mat_inv M = k' *\<^sub>s\<^sub>m (a', b', cnj b', cnj a') \ mat_det (a', b', cnj b', cnj a') \ 0" using unitary11_inv [of k M a b] by auto thus ?thesis using unitary11_gen_iff[of "mat_inv M"] by auto qed lemma unitary11_gen_comp: assumes "unitary11_gen M1" and "mat_det M1 \ 0" and "unitary11_gen M2" and "mat_det M2 \ 0" shows "unitary11_gen (M1 *\<^sub>m\<^sub>m M2)" proof- from assms obtain k1 k2 a1 a2 b1 b2 where "k1 \ 0 \ mat_det (a1, b1, cnj b1, cnj a1) \ 0 \ M1 = k1 *\<^sub>s\<^sub>m (a1, b1, cnj b1, cnj a1)" "k2 \ 0 \ mat_det (a2, b2, cnj b2, cnj a2) \ 0 \ M2 = k2 *\<^sub>s\<^sub>m (a2, b2, cnj b2, cnj a2)" using unitary11_gen_iff[of M1] unitary11_gen_iff[of M2] by blast then obtain k a b where "k \ 0 \ M1 *\<^sub>m\<^sub>m M2 = k *\<^sub>s\<^sub>m (a, b, cnj b, cnj a) \ mat_det (a, b, cnj b, cnj a) \ 0" using unitary11_comp[of k1 M1 a1 b1 k2 M2 a2 b2] by blast thus ?thesis using unitary11_gen_iff[of "M1 *\<^sub>m\<^sub>m M2"] by blast qed text \Classification into orientation-preserving and orientation-reversing matrices\ lemma unitary11_sgn_det_orientation: assumes "k \ 0" and "mat_det (a, b, cnj b, cnj a) \ 0" and "M = k *\<^sub>s\<^sub>m (a, b, cnj b, cnj a)" shows "\ k'. sgn k' = sgn (Re (mat_det (a, b, cnj b, cnj a))) \ congruence M (1, 0, 0, -1) = cor k' *\<^sub>s\<^sub>m (1, 0, 0, -1)" proof- let ?x = "cnj k * cnj a * (k * a) - (cnj k * b * (k * cnj b))" have *: "?x = k * cnj k * (a * cnj a - b * cnj b)" by (auto simp add: field_simps) hence "is_real ?x" by auto hence "cor (Re ?x) = ?x" by (rule complex_of_real_Re) moreover have "sgn (Re ?x) = sgn (Re (a * cnj a - b * cnj b))" proof- have *: "Re ?x = (cmod k)\<^sup>2 * Re (a * cnj a - b * cnj b)" by (subst *, subst complex_mult_cnj_cmod, subst Re_mult_real) (metis Im_complex_of_real, metis Re_complex_of_real) show ?thesis using \k \ 0\ by (subst *) (simp add: sgn_mult) qed ultimately show ?thesis using assms(3) by (rule_tac x="Re ?x" in exI) (auto simp add: mat_adj_def mat_cnj_def) qed lemma unitary11_sgn_det: assumes "k \ 0" and "mat_det (a, b, cnj b, cnj a) \ 0" and "M = k *\<^sub>s\<^sub>m (a, b, cnj b, cnj a)" and "M = (A, B, C, D)" shows "sgn (Re (mat_det (a, b, cnj b, cnj a))) = (if b = 0 then 1 else sgn (Re ((A*D)/(B*C)) - 1))" proof (cases "b = 0") case True thus ?thesis using assms by (simp only: mat_det.simps, subst complex_mult_cnj_cmod, subst minus_complex.sel, subst Re_complex_of_real, simp) next case False from assms have *: "A = k * a" "B = k * b" "C = k * cnj b" "D = k * cnj a" by auto hence *: "(A*D)/(B*C) = (a*cnj a)/(b*cnj b)" using \k \ 0\ by simp show ?thesis using \b \ 0\ apply (subst *, subst Re_divide_real, simp, simp) apply (simp only: mat_det.simps) apply (subst complex_mult_cnj_cmod)+ apply ((subst Re_complex_of_real)+, subst minus_complex.sel, (subst Re_complex_of_real)+, simp add: field_simps sgn_if) done qed lemma unitary11_orientation: assumes "unitary11_gen M" and "M = (A, B, C, D)" shows "\ k'. sgn k' = sgn (if B = 0 then 1 else sgn (Re ((A*D)/(B*C)) - 1)) \ congruence M (1, 0, 0, -1) = cor k' *\<^sub>s\<^sub>m (1, 0, 0, -1)" proof- from \unitary11_gen M\ obtain k a b where *: "k \ 0" "mat_det (a, b, cnj b, cnj a) \ 0" "M = k*\<^sub>s\<^sub>m (a, b, cnj b, cnj a)" using unitary11_gen_iff[of M] by auto moreover have "b = 0 \ B = 0" using \M = (A, B, C, D)\ * by auto ultimately show ?thesis using unitary11_sgn_det_orientation[OF *] unitary11_sgn_det[OF * \M = (A, B, C, D)\] by auto qed lemma unitary11_sgn_det_orientation': assumes "congruence M (1, 0, 0, -1) = cor k' *\<^sub>s\<^sub>m (1, 0, 0, -1)" and "k' \ 0" shows "\ a b k. k \ 0 \ M = k *\<^sub>s\<^sub>m (a, b, cnj b, cnj a) \ sgn k' = sgn (Re (mat_det (a, b, cnj b, cnj a)))" proof- obtain a b k where "k \ 0" "mat_det (a, b, cnj b, cnj a) \ 0" "M = k *\<^sub>s\<^sub>m (a, b, cnj b, cnj a)" using assms using unitary11_gen_iff[of M] unfolding unitary11_gen_def by auto moreover have "sgn k' = sgn (Re (mat_det (a, b, cnj b, cnj a)))" proof- let ?x = "cnj k * cnj a * (k * a) - (cnj k * b * (k * cnj b))" have *: "?x = k * cnj k * (a * cnj a - b * cnj b)" by (auto simp add: field_simps) hence "is_real ?x" by auto hence "cor (Re ?x) = ?x" by (rule complex_of_real_Re) have **: "sgn (Re ?x) = sgn (Re (a * cnj a - b * cnj b))" proof- have *: "Re ?x = (cmod k)\<^sup>2 * Re (a * cnj a - b * cnj b)" by (subst *, subst complex_mult_cnj_cmod, subst Re_mult_real) (metis Im_complex_of_real, metis Re_complex_of_real) show ?thesis using \k \ 0\ by (subst *) (simp add: sgn_mult) qed moreover have "?x = cor k'" using \M = k *\<^sub>s\<^sub>m (a, b, cnj b, cnj a)\ assms by (simp add: mat_adj_def mat_cnj_def) hence "sgn (Re ?x) = sgn k'" using \cor (Re ?x) = ?x\ unfolding complex_of_real_def by simp ultimately show ?thesis by simp qed ultimately show ?thesis by (rule_tac x="a" in exI, rule_tac x="b" in exI, rule_tac x="k" in exI) simp qed end diff --git a/thys/Complex_Geometry/Unitary_Matrices.thy b/thys/Complex_Geometry/Unitary_Matrices.thy --- a/thys/Complex_Geometry/Unitary_Matrices.thy +++ b/thys/Complex_Geometry/Unitary_Matrices.thy @@ -1,330 +1,330 @@ (* -------------------------------------------------------------------------- *) subsection \Generalized Unitary Matrices\ (* -------------------------------------------------------------------------- *) theory Unitary_Matrices imports Matrices More_Complex begin text \In this section (generalized) $2\times 2$ unitary matrices are introduced.\ text \Unitary matrices\ definition unitary where "unitary M \ mat_adj M *\<^sub>m\<^sub>m M = eye" text \Generalized unitary matrices\ definition unitary_gen where "unitary_gen M \ (\ k::complex. k \ 0 \ mat_adj M *\<^sub>m\<^sub>m M = k *\<^sub>s\<^sub>m eye)" text \Scalar can be always be a positive real\ lemma unitary_gen_real: assumes "unitary_gen M" shows "(\ k::real. k > 0 \ mat_adj M *\<^sub>m\<^sub>m M = cor k *\<^sub>s\<^sub>m eye)" proof- obtain k where *: "mat_adj M *\<^sub>m\<^sub>m M = k *\<^sub>s\<^sub>m eye" "k \ 0" using assms by (auto simp add: unitary_gen_def) obtain a b c d where "M = (a, b, c, d)" by (cases M) auto hence "k = cor ((cmod a)\<^sup>2) + cor ((cmod c)\<^sup>2)" using * by (subst complex_mult_cnj_cmod[symmetric])+ (auto simp add: mat_adj_def mat_cnj_def) hence "is_real k \ Re k > 0" using \k \ 0\ by (smt add_cancel_left_left arg_0_iff arg_complex_of_real_positive not_sum_power2_lt_zero of_real_0 plus_complex.simps(1) plus_complex.simps(2)) thus ?thesis using * by (rule_tac x="Re k" in exI) simp qed text \Generalized unitary matrices can be factored into a product of a unitary matrix and a real positive scalar multiple of the identity matrix\ lemma unitary_gen_unitary: shows "unitary_gen M \ (\ k M'. k > 0 \ unitary M' \ M = (cor k *\<^sub>s\<^sub>m eye) *\<^sub>m\<^sub>m M')" (is "?lhs = ?rhs") proof assume ?lhs then obtain k where *: "k>0" "mat_adj M *\<^sub>m\<^sub>m M = cor k *\<^sub>s\<^sub>m eye" using unitary_gen_real[of M] by auto let ?k' = "cor (sqrt k)" have "?k' * cnj ?k' = cor k" using \k > 0\ by simp moreover have "Re ?k' > 0" "is_real ?k'" "?k' \ 0" using \k > 0\ by auto ultimately show ?rhs using * mat_eye_l unfolding unitary_gen_def unitary_def by (rule_tac x="Re ?k'" in exI) (rule_tac x="(1/?k')*\<^sub>s\<^sub>mM" in exI, simp add: mult_sm_mm[symmetric]) next assume ?rhs then obtain k M' where "k > 0" "unitary M'" "M = (cor k *\<^sub>s\<^sub>m eye) *\<^sub>m\<^sub>m M'" by blast hence "M = cor k *\<^sub>s\<^sub>m M'" using mult_sm_mm[of "cor k" eye M'] mat_eye_l by simp thus ?lhs using \unitary M'\ \k > 0\ by (simp add: unitary_gen_def unitary_def) qed text \When they represent Möbius transformations, eneralized unitary matrices fix the imaginary unit circle. Therefore, they fix a Hermitean form with (2, 0) signature (two positive and no negative diagonal elements).\ lemma unitary_gen_iff': shows "unitary_gen M \ (\ k::complex. k \ 0 \ congruence M (1, 0, 0, 1) = k *\<^sub>s\<^sub>m (1, 0, 0, 1))" unfolding unitary_gen_def using mat_eye_r by (auto simp add: mult.assoc) text \Unitary matrices are special cases of general unitary matrices\ lemma unitary_unitary_gen [simp]: assumes "unitary M" shows "unitary_gen M" using assms unfolding unitary_gen_def unitary_def by auto text \Generalized unitary matrices are regular\ lemma unitary_gen_regular: assumes "unitary_gen M" shows "mat_det M \ 0" proof- from assms obtain k where "k \ 0" "mat_adj M *\<^sub>m\<^sub>m M = k *\<^sub>s\<^sub>m eye" unfolding unitary_gen_def by auto hence "mat_det (mat_adj M *\<^sub>m\<^sub>m M) \ 0" by simp thus ?thesis by (simp add: mat_det_adj) qed lemmas unitary_regular = unitary_gen_regular[OF unitary_unitary_gen] (* -------------------------------------------------------------------------- *) subsubsection \Group properties\ (* -------------------------------------------------------------------------- *) text \Generalized $2\times 2$ unitary matrices form a group under multiplication (usually denoted by $GU(2, \mathbb{C})$). The group is closed under non-zero complex scalar multiplication. Since these matrices are always regular, they form a subgroup of general linear group (usually denoted by $GL(2, \mathbb{C})$) of all regular matrices.\ lemma unitary_gen_scale [simp]: assumes "unitary_gen M" and "k \ 0" shows "unitary_gen (k *\<^sub>s\<^sub>m M)" using assms unfolding unitary_gen_def by auto lemma unitary_comp: assumes "unitary M1" and "unitary M2" shows "unitary (M1 *\<^sub>m\<^sub>m M2)" using assms unfolding unitary_def by (metis mat_adj_mult_mm mat_eye_l mult_mm_assoc) lemma unitary_gen_comp: assumes "unitary_gen M1" and "unitary_gen M2" shows "unitary_gen (M1 *\<^sub>m\<^sub>m M2)" proof- obtain k1 k2 where *: "k1 * k2 \ 0" "mat_adj M1 *\<^sub>m\<^sub>m M1 = k1 *\<^sub>s\<^sub>m eye" "mat_adj M2 *\<^sub>m\<^sub>m M2 = k2 *\<^sub>s\<^sub>m eye" using assms unfolding unitary_gen_def by auto have "mat_adj M2 *\<^sub>m\<^sub>m mat_adj M1 *\<^sub>m\<^sub>m (M1 *\<^sub>m\<^sub>m M2) = mat_adj M2 *\<^sub>m\<^sub>m (mat_adj M1 *\<^sub>m\<^sub>m M1) *\<^sub>m\<^sub>m M2" by (auto simp add: mult_mm_assoc) also have "... = mat_adj M2 *\<^sub>m\<^sub>m ((k1 *\<^sub>s\<^sub>m eye) *\<^sub>m\<^sub>m M2)" using * by (auto simp add: mult_mm_assoc) also have "... = mat_adj M2 *\<^sub>m\<^sub>m (k1 *\<^sub>s\<^sub>m M2)" using mult_sm_eye_mm[of k1 M2] by (simp del: eye_def) also have "... = k1 *\<^sub>s\<^sub>m (k2 *\<^sub>s\<^sub>m eye)" using * by auto finally show ?thesis using * unfolding unitary_gen_def by (rule_tac x="k1*k2" in exI, simp del: eye_def) qed lemma unitary_adj_eq_inv: shows "unitary M \ mat_det M \ 0 \ mat_adj M = mat_inv M" using unitary_regular[of M] mult_mm_inv_r[of M "mat_adj M" eye] mat_eye_l[of "mat_inv M"] mat_inv_l[of M] unfolding unitary_def by - (rule, simp_all) lemma unitary_inv: assumes "unitary M" shows "unitary (mat_inv M)" using assms unfolding unitary_adj_eq_inv using mat_adj_inv[of M] mat_det_inv[of M] by simp lemma unitary_gen_inv: assumes "unitary_gen M" shows "unitary_gen (mat_inv M)" proof- obtain k M' where "0 < k" "unitary M'" "M = cor k *\<^sub>s\<^sub>m eye *\<^sub>m\<^sub>m M'" using unitary_gen_unitary[of M] assms by blast hence "mat_inv M = cor (1/k) *\<^sub>s\<^sub>m mat_inv M'" by (metis mat_inv_mult_sm mult_sm_eye_mm norm_not_less_zero of_real_1 of_real_divide of_real_eq_0_iff sgn_1_neg sgn_greater sgn_if sgn_pos sgn_sgn) thus ?thesis using \k > 0\ \unitary M'\ by (subst unitary_gen_unitary[of "mat_inv M"]) (rule_tac x="1/k" in exI, rule_tac x="mat_inv M'" in exI, metis divide_pos_pos mult_sm_eye_mm unitary_inv zero_less_one) qed (* -------------------------------------------------------------------------- *) subsubsection \The characterization in terms of matrix elements\ (* -------------------------------------------------------------------------- *) text \Special matrices are those having the determinant equal to 1. We first give their characterization.\ lemma unitary_special: assumes "unitary M" and "mat_det M = 1" shows "\ a b. M = (a, b, -cnj b, cnj a)" proof- have "mat_adj M = mat_inv M" using assms mult_mm_inv_r[of M "mat_adj M" "eye"] mat_eye_r mat_eye_l by (simp add: unitary_def) thus ?thesis using \mat_det M = 1\ by (cases M) (auto simp add: mat_adj_def mat_cnj_def) qed lemma unitary_gen_special: assumes "unitary_gen M" and "mat_det M = 1" shows "\ a b. M = (a, b, -cnj b, cnj a)" proof- from assms obtain k where *: "k \ 0" "mat_adj M *\<^sub>m\<^sub>m M = k *\<^sub>s\<^sub>m eye" unfolding unitary_gen_def by auto hence "mat_det (mat_adj M *\<^sub>m\<^sub>m M) = k*k" by simp hence "k*k = 1" using assms(2) by (simp add: mat_det_adj) hence "k = 1 \ k = -1" using square_eq_1_iff[of k] by simp moreover have "mat_adj M = k *\<^sub>s\<^sub>m mat_inv M" using * using assms mult_mm_inv_r[of M "mat_adj M" "k *\<^sub>s\<^sub>m eye"] mat_eye_r mat_eye_l by simp (metis mult_sm_eye_mm *(2)) moreover obtain a b c d where "M = (a, b, c, d)" by (cases M) auto ultimately have "M = (a, b, -cnj b, cnj a) \ M = (a, b, cnj b, -cnj a)" using assms(2) by (auto simp add: mat_adj_def mat_cnj_def) moreover have "Re (- (cor (cmod a))\<^sup>2 - (cor (cmod b))\<^sup>2) < 1" by (smt cmod_square complex_norm_square minus_complex.simps(1) of_real_power realpow_square_minus_le uminus_complex.simps(1)) hence "- (cor (cmod a))\<^sup>2 - (cor (cmod b))\<^sup>2 \ 1" by force hence "M \ (a, b, cnj b, -cnj a)" using \mat_det M = 1\ complex_mult_cnj_cmod[of a] complex_mult_cnj_cmod[of b] by auto ultimately show ?thesis by auto qed text \A characterization of all generalized unitary matrices\ lemma unitary_gen_iff: shows "unitary_gen M \ (\ a b k. k \ 0 \ mat_det (a, b, -cnj b, cnj a) \ 0 \ M = k *\<^sub>s\<^sub>m (a, b, -cnj b, cnj a))" (is "?lhs = ?rhs") proof assume ?lhs obtain d where *: "d*d = mat_det M" using ex_complex_sqrt by auto hence "d \ 0" using unitary_gen_regular[OF \unitary_gen M\] by auto from \unitary_gen M\ obtain k where "k \ 0" "mat_adj M *\<^sub>m\<^sub>m M = k *\<^sub>s\<^sub>m eye" unfolding unitary_gen_def by auto hence "mat_adj ((1/d)*\<^sub>s\<^sub>mM) *\<^sub>m\<^sub>m ((1/d)*\<^sub>s\<^sub>mM) = (k / (d*cnj d)) *\<^sub>s\<^sub>m eye" by simp obtain a b where "(a, b, - cnj b, cnj a) = (1 / d) *\<^sub>s\<^sub>m M" using unitary_gen_special[of "(1 / d) *\<^sub>s\<^sub>m M"] \unitary_gen M\ * unitary_gen_regular[of M] \d \ 0\ by force moreover hence "mat_det (a, b, - cnj b, cnj a) \ 0" using unitary_gen_regular[OF \unitary_gen M\] \d \ 0\ by auto ultimately show ?rhs apply (rule_tac x="a" in exI, rule_tac x="b" in exI, rule_tac x="d" in exI) using mult_sm_inv_l[of "1/d" M] by (auto simp add: field_simps) next assume ?rhs then obtain a b k where "k \ 0 \ mat_det (a, b, - cnj b, cnj a) \ 0 \ M = k *\<^sub>s\<^sub>m (a, b, - cnj b, cnj a)" by auto thus ?lhs unfolding unitary_gen_def apply (auto simp add: mat_adj_def mat_cnj_def) using mult_eq_0_iff[of "cnj k * k" "cnj a * a + cnj b * b"] by (auto simp add: field_simps) qed text \A characterization of unitary matrices\ lemma unitary_iff: shows "unitary M \ (\ a b k. (cmod a)\<^sup>2 + (cmod b)\<^sup>2 \ 0 \ (cmod k)\<^sup>2 = 1 / ((cmod a)\<^sup>2 + (cmod b)\<^sup>2) \ M = k *\<^sub>s\<^sub>m (a, b, -cnj b, cnj a))" (is "?lhs = ?rhs") proof assume ?lhs obtain k a b where *: "M = k *\<^sub>s\<^sub>m (a, b, -cnj b, cnj a)" "k \ 0" "mat_det (a, b, -cnj b, cnj a) \ 0" using unitary_gen_iff unitary_unitary_gen[OF \unitary M\] by auto have md: "mat_det (a, b, -cnj b, cnj a) = cor ((cmod a)\<^sup>2 + (cmod b)\<^sup>2)" by (auto simp add: complex_mult_cnj_cmod) have "k * cnj k * mat_det (a, b, -cnj b, cnj a) = 1" using \unitary M\ * unfolding unitary_def by (auto simp add: mat_adj_def mat_cnj_def field_simps) hence "(cmod k)\<^sup>2 * ((cmod a)\<^sup>2 + (cmod b)\<^sup>2) = 1" - by (subst (asm) complex_mult_cnj_cmod, subst (asm) md, subst (asm) cor_mult[symmetric]) (metis of_real_1 of_real_eq_iff) + by (metis (mono_tags, lifting) complex_norm_square md of_real_1 of_real_eq_iff of_real_mult) thus ?rhs using * mat_eye_l apply (rule_tac x="a" in exI, rule_tac x="b" in exI, rule_tac x="k" in exI) apply (auto simp add: complex_mult_cnj_cmod) by (metis \(cmod k)\<^sup>2 * ((cmod a)\<^sup>2 + (cmod b)\<^sup>2) = 1\ mult_eq_0_iff nonzero_eq_divide_eq zero_neq_one) next assume ?rhs then obtain a b k where *: "(cmod a)\<^sup>2 + (cmod b)\<^sup>2 \ 0" "(cmod k)\<^sup>2 = 1 / ((cmod a)\<^sup>2 + (cmod b)\<^sup>2)" "M = k *\<^sub>s\<^sub>m (a, b, -cnj b, cnj a)" by auto have "(k * cnj k) * (a * cnj a) + (k * cnj k) * (b * cnj b) = 1" apply (subst complex_mult_cnj_cmod)+ using *(1-2) by (metis (no_types, lifting) distrib_left nonzero_eq_divide_eq of_real_1 of_real_add of_real_divide of_real_eq_0_iff) thus ?lhs using * unfolding unitary_def by (simp add: mat_adj_def mat_cnj_def field_simps) qed end