diff --git a/thys/Complex_Geometry/Angles.thy b/thys/Complex_Geometry/Angles.thy --- a/thys/Complex_Geometry/Angles.thy +++ b/thys/Complex_Geometry/Angles.thy @@ -1,510 +1,510 @@ (* ---------------------------------------------------------------------------- *) subsection \Angle between two vectors\ (* ---------------------------------------------------------------------------- *) text \In this section we introduce different measures of angle between two vectors (represented by complex numbers).\ theory Angles imports More_Transcendental Canonical_Angle More_Complex begin (* ---------------------------------------------------------------------------- *) subsubsection \Oriented angle\ (* ---------------------------------------------------------------------------- *) text \Oriented angle between two vectors (it is always in the interval $(-\pi, \pi]$).\ definition ang_vec ("\") where - [simp]: "\ z1 z2 \ \arg z2 - arg z1\" + [simp]: "\ z1 z2 \ \Arg z2 - Arg z1\" lemma ang_vec_bounded: shows "-pi < \ z1 z2 \ \ z1 z2 \ pi" by (simp add: canon_ang(1) canon_ang(2)) lemma ang_vec_sym: assumes "\ z1 z2 \ pi" shows "\ z1 z2 = - \ z2 z1" using assms unfolding ang_vec_def - using canon_ang_uminus[of "arg z2 - arg z1"] + using canon_ang_uminus[of "Arg z2 - Arg z1"] by simp lemma ang_vec_sym_pi: assumes "\ z1 z2 = pi" shows "\ z1 z2 = \ z2 z1" using assms unfolding ang_vec_def - using canon_ang_uminus_pi[of "arg z2 - arg z1"] + using canon_ang_uminus_pi[of "Arg z2 - Arg z1"] by simp lemma ang_vec_plus_pi1: assumes "\ z1 z2 > 0" shows "\\ z1 z2 + pi\ = \ z1 z2 - pi" proof (rule canon_ang_eqI) show "\ x::int. \ z1 z2 - pi - (\ z1 z2 + pi) = 2 * real_of_int x * pi" by (rule_tac x="-1" in exI) auto next show "- pi < \ z1 z2 - pi \ \ z1 z2 - pi \ pi" using assms unfolding ang_vec_def - using canon_ang(1)[of "arg z2 - arg z1"] canon_ang(2)[of "arg z2 - arg z1"] + using canon_ang(1)[of "Arg z2 - Arg z1"] canon_ang(2)[of "Arg z2 - Arg z1"] by auto qed lemma ang_vec_plus_pi2: assumes "\ z1 z2 \ 0" shows "\\ z1 z2 + pi\ = \ z1 z2 + pi" proof (rule canon_ang_id) show "- pi < \ z1 z2 + pi \ \ z1 z2 + pi \ pi" using assms unfolding ang_vec_def - using canon_ang(1)[of "arg z2 - arg z1"] canon_ang(2)[of "arg z2 - arg z1"] + using canon_ang(1)[of "Arg z2 - Arg z1"] canon_ang(2)[of "Arg z2 - Arg z1"] by auto qed lemma ang_vec_opposite1: assumes "z1 \ 0" shows "\ (-z1) z2 = \\ z1 z2 - pi\" proof- - have "\ (-z1) z2 = \arg z2 - (arg z1 + pi)\" + have "\ (-z1) z2 = \Arg z2 - (Arg z1 + pi)\" unfolding ang_vec_def using arg_uminus[OF assms] using canon_ang_arg[of z2, symmetric] - using canon_ang_diff[of "arg z2" "arg z1 + pi", symmetric] + using canon_ang_diff[of "Arg z2" "Arg z1 + pi", symmetric] by simp moreover - have "\\ z1 z2 - pi\ = \arg z2 - arg z1 - pi\" + have "\\ z1 z2 - pi\ = \Arg z2 - Arg z1 - pi\" using canon_ang_id[of pi, symmetric] - using canon_ang_diff[of "arg z2 - arg z1" "pi", symmetric] + using canon_ang_diff[of "Arg z2 - Arg z1" "pi", symmetric] by simp_all ultimately show ?thesis by (simp add: field_simps) qed lemma ang_vec_opposite2: assumes "z2 \ 0" shows "\ z1 (-z2) = \\ z1 z2 + pi\" unfolding ang_vec_def using arg_mult[of "-1" "z2"] assms using arg_complex_of_real_negative[of "-1"] - using canon_ang_diff[of "arg (-1) + arg z2" "arg z1", symmetric] - using canon_ang_sum[of "arg z2 - arg z1" "pi", symmetric] + using canon_ang_diff[of "Arg (-1) + Arg z2" "Arg z1", symmetric] + using canon_ang_sum[of "Arg z2 - Arg z1" "pi", symmetric] using canon_ang_id[of pi] canon_ang_arg[of z1] by (auto simp: algebra_simps) lemma ang_vec_opposite_opposite: assumes "z1 \ 0" and "z2 \ 0" shows "\ (-z1) (-z2) = \ z1 z2" proof- have "\ (-z1) (-z2) = \\\ z1 z2 + pi\ - \pi\\" using ang_vec_opposite1[OF assms(1)] using ang_vec_opposite2[OF assms(2)] using canon_ang_id[of pi, symmetric] by simp_all also have "... = \\ z1 z2\" by (subst canon_ang_diff[symmetric], simp) finally show ?thesis by (metis ang_vec_def canon_ang(1) canon_ang(2) canon_ang_id) qed lemma ang_vec_opposite_opposite': assumes "z1 \ z" and "z2 \ z" shows "\ (z - z1) (z - z2) = \ (z1 - z) (z2 - z)" using ang_vec_opposite_opposite[of "z - z1" "z - z2"] assms by (simp add: field_simps del: ang_vec_def) text \Cosine, scalar product and the law of cosines\ lemma cos_cmod_scalprod: shows "cmod z1 * cmod z2 * (cos (\ z1 z2)) = Re (scalprod z1 z2)" proof (cases "z1 = 0 \ z2 = 0") case True thus ?thesis by auto next case False thus ?thesis by (simp add: cos_diff cos_arg sin_arg field_simps) qed lemma cos0_scalprod0: assumes "z1 \ 0" and "z2 \ 0" shows "cos (\ z1 z2) = 0 \ scalprod z1 z2 = 0" using assms using cnj_mix_real[of z1 z2] using cos_cmod_scalprod[of z1 z2] by (auto simp add: complex_eq_if_Re_eq) lemma ortho_scalprod0: assumes "z1 \ 0" and "z2 \ 0" shows "\ z1 z2 = pi/2 \ \ z1 z2 = -pi/2 \ scalprod z1 z2 = 0" using cos0_scalprod0[OF assms] using ang_vec_bounded[of z1 z2] using cos_0_iff_canon[of "\ z1 z2"] by (metis cos_minus cos_pi_half divide_minus_left) lemma law_of_cosines: shows "(cdist B C)\<^sup>2 = (cdist A C)\<^sup>2 + (cdist A B)\<^sup>2 - 2*(cdist A C)*(cdist A B)*(cos (\ (C-A) (B-A)))" proof- let ?a = "C-B" and ?b = "C-A" and ?c = "B-A" have "?a = ?b - ?c" by simp hence "(cmod ?a)\<^sup>2 = (cmod (?b - ?c))\<^sup>2" by metis also have "... = Re (scalprod (?b-?c) (?b-?c))" by (simp add: cmod_square) also have "... = (cmod ?b)\<^sup>2 + (cmod ?c)\<^sup>2 - 2*Re (scalprod ?b ?c)" by (simp add: cmod_square field_simps) finally show ?thesis using cos_cmod_scalprod[of ?b ?c] by simp qed (* ---------------------------------------------------------------------------- *) subsubsection \Unoriented angle\ (* ---------------------------------------------------------------------------- *) text \Convex unoriented angle between two vectors (it is always in the interval $[0, pi]$).\ definition ang_vec_c ("\c") where [simp]:"\c z1 z2 \ abs (\ z1 z2)" lemma ang_vec_c_sym: shows "\c z1 z2 = \c z2 z1" unfolding ang_vec_c_def using ang_vec_sym_pi[of z1 z2] ang_vec_sym[of z1 z2] by (cases "\ z1 z2 = pi") auto lemma ang_vec_c_bounded: "0 \ \c z1 z2 \ \c z1 z2 \ pi" - using canon_ang(1)[of "arg z2 - arg z1"] canon_ang(2)[of "arg z2 - arg z1"] + using canon_ang(1)[of "Arg z2 - Arg z1"] canon_ang(2)[of "Arg z2 - Arg z1"] by auto text \Cosine and scalar product\ lemma cos_c_: "cos (\c z1 z2) = cos (\ z1 z2)" unfolding ang_vec_c_def by (smt cos_minus) lemma ortho_c_scalprod0: assumes "z1 \ 0" and "z2 \ 0" shows "\c z1 z2 = pi/2 \ scalprod z1 z2 = 0" proof- have "\ z1 z2 = pi / 2 \ \ z1 z2 = - pi / 2 \ \c z1 z2 = pi/2" unfolding ang_vec_c_def using arctan by force thus ?thesis using ortho_scalprod0[OF assms] by simp qed (* ---------------------------------------------------------------------------- *) subsubsection \Acute angle\ (* ---------------------------------------------------------------------------- *) text \Acute or right angle (non-obtuse) between two vectors (it is always in the interval $[0, \frac{\pi}{2}$]). We will use this to measure angle between two circles, since it can always be acute (or right).\ definition acute_ang where [simp]: "acute_ang \ = (if \ > pi / 2 then pi - \ else \)" definition ang_vec_a ("\a") where [simp]: "\a z1 z2 \ acute_ang (\c z1 z2)" lemma ang_vec_a_sym: "\a z1 z2 = \a z2 z1" unfolding ang_vec_a_def using ang_vec_c_sym by auto lemma ang_vec_a_opposite2: "\a z1 z2 = \a z1 (-z2)" proof(cases "z2 = 0") case True thus ?thesis by (metis minus_zero) next case False thus ?thesis proof(cases "\ z1 z2 < -pi / 2") case True hence "\ z1 z2 < 0" using pi_not_less_zero by linarith have "\a z1 z2 = pi + \ z1 z2" using True \\ z1 z2 < 0\ unfolding ang_vec_a_def ang_vec_c_def ang_vec_a_def abs_real_def by auto moreover have "\a z1 (-z2) = pi + \ z1 z2" unfolding ang_vec_a_def ang_vec_c_def abs_real_def - using canon_ang(1)[of "arg z2 - arg z1"] canon_ang(2)[of "arg z2 - arg z1"] + using canon_ang(1)[of "Arg z2 - Arg z1"] canon_ang(2)[of "Arg z2 - Arg z1"] using ang_vec_plus_pi2[of z1 z2] True \\ z1 z2 < 0\ \z2 \ 0\ using ang_vec_opposite2[of z2 z1] by auto ultimately show ?thesis by auto next case False show ?thesis proof (cases "\ z1 z2 \ 0") case True have "\a z1 z2 = - \ z1 z2" using \\ \ z1 z2 < - pi / 2\ True unfolding ang_vec_a_def ang_vec_c_def ang_vec_a_def abs_real_def by auto moreover have "\a z1 (-z2) = - \ z1 z2" using \\ \ z1 z2 < - pi / 2\ True unfolding ang_vec_a_def ang_vec_c_def abs_real_def using ang_vec_plus_pi2[of z1 z2] - using canon_ang(1)[of "arg z2 - arg z1"] canon_ang(2)[of "arg z2 - arg z1"] + using canon_ang(1)[of "Arg z2 - Arg z1"] canon_ang(2)[of "Arg z2 - Arg z1"] using \z2 \ 0\ ang_vec_opposite2[of z2 z1] by auto ultimately show ?thesis by simp next case False show ?thesis proof (cases "\ z1 z2 < pi / 2") case True have "\a z1 z2 = \ z1 z2" using \\ \ z1 z2 \ 0\ True unfolding ang_vec_a_def ang_vec_c_def ang_vec_a_def abs_real_def by auto moreover have "\a z1 (-z2) = \ z1 z2" using \\ \ z1 z2 \ 0\ True unfolding ang_vec_a_def ang_vec_c_def abs_real_def using ang_vec_plus_pi1[of z1 z2] - using canon_ang(1)[of "arg z2 - arg z1"] canon_ang(2)[of "arg z2 - arg z1"] + using canon_ang(1)[of "Arg z2 - Arg z1"] canon_ang(2)[of "Arg z2 - Arg z1"] using \z2 \ 0\ ang_vec_opposite2[of z2 z1] by auto ultimately show ?thesis by simp next case False have "\ z1 z2 > 0" using False by (metis less_linear less_trans pi_half_gt_zero) have "\a z1 z2 = pi - \ z1 z2" using False \\ z1 z2 > 0\ unfolding ang_vec_a_def ang_vec_c_def ang_vec_a_def abs_real_def by auto moreover have "\a z1 (-z2) = pi - \ z1 z2" unfolding ang_vec_a_def ang_vec_c_def abs_real_def using False \\ z1 z2 > 0\ using ang_vec_plus_pi1[of z1 z2] - using canon_ang(1)[of "arg z2 - arg z1"] canon_ang(2)[of "arg z2 - arg z1"] + using canon_ang(1)[of "Arg z2 - Arg z1"] canon_ang(2)[of "Arg z2 - Arg z1"] using \z2 \ 0\ ang_vec_opposite2[of z2 z1] by auto ultimately show ?thesis by auto qed qed qed qed lemma ang_vec_a_opposite1: shows "\a z1 z2 = \a (-z1) z2" using ang_vec_a_sym[of "-z1" z2] ang_vec_a_opposite2[of z2 z1] ang_vec_a_sym[of z2 z1] by auto lemma ang_vec_a_scale1: assumes "k \ 0" shows "\a (cor k * z1) z2 = \a z1 z2" proof (cases "k > 0") case True thus ?thesis unfolding ang_vec_a_def ang_vec_c_def ang_vec_def using arg_mult_real_positive[of k z1] by auto next case False hence "k < 0" using assms by auto thus ?thesis using arg_mult_real_negative[of k z1] using ang_vec_a_opposite1[of z1 z2] unfolding ang_vec_a_def ang_vec_c_def ang_vec_def by simp qed lemma ang_vec_a_scale2: assumes "k \ 0" shows "\a z1 (cor k * z2) = \a z1 z2" using ang_vec_a_sym[of z1 "complex_of_real k * z2"] using ang_vec_a_scale1[OF assms, of z2 z1] using ang_vec_a_sym[of z1 z2] by auto lemma ang_vec_a_scale: assumes "k1 \ 0" and "k2 \ 0" shows "\a (cor k1 * z1) (cor k2 * z2) = \a z1 z2" using ang_vec_a_scale1[OF assms(1)] ang_vec_a_scale2[OF assms(2)] by auto lemma ang_a_cnj_cnj: shows "\a z1 z2 = \a (cnj z1) (cnj z2)" unfolding ang_vec_a_def ang_vec_c_def ang_vec_def -proof(cases "arg z1 \ pi \ arg z2 \ pi") +proof(cases "Arg z1 \ pi \ Arg z2 \ pi") case True - thus "acute_ang \\arg z2 - arg z1\\ = acute_ang \\arg (cnj z2) - arg (cnj z1)\\" + thus "acute_ang \\Arg z2 - Arg z1\\ = acute_ang \\Arg (cnj z2) - Arg (cnj z1)\\" using arg_cnj_not_pi[of z1] arg_cnj_not_pi[of z2] apply (auto simp del:acute_ang_def) - proof(cases "\arg z2 - arg z1\ = pi") + proof(cases "\Arg z2 - Arg z1\ = pi") case True - thus "acute_ang \\arg z2 - arg z1\\ = acute_ang \\arg z1 - arg z2\\" - using canon_ang_uminus_pi[of "arg z2 - arg z1"] + thus "acute_ang \\Arg z2 - Arg z1\\ = acute_ang \\Arg z1 - Arg z2\\" + using canon_ang_uminus_pi[of "Arg z2 - Arg z1"] by (auto simp add:field_simps) next case False - thus "acute_ang \\arg z2 - arg z1\\ = acute_ang \\arg z1 - arg z2\\" - using canon_ang_uminus[of "arg z2 - arg z1"] + thus "acute_ang \\Arg z2 - Arg z1\\ = acute_ang \\Arg z1 - Arg z2\\" + using canon_ang_uminus[of "Arg z2 - Arg z1"] by (auto simp add:field_simps) qed next case False - thus "acute_ang \\arg z2 - arg z1\\ = acute_ang \\arg (cnj z2) - arg (cnj z1)\\" - proof(cases "arg z1 = pi") + thus "acute_ang \\Arg z2 - Arg z1\\ = acute_ang \\Arg (cnj z2) - Arg (cnj z1)\\" + proof(cases "Arg z1 = pi") case False - hence "arg z2 = pi" - using \ \ (arg z1 \ pi \ arg z2 \ pi)\ + hence "Arg z2 = pi" + using \ \ (Arg z1 \ pi \ Arg z2 \ pi)\ by auto thus ?thesis using False using arg_cnj_not_pi[of z1] arg_cnj_pi[of z2] apply (auto simp del:acute_ang_def) - proof(cases "arg z1 > 0") + proof(cases "Arg z1 > 0") case True - hence "-arg z1 \ 0" + hence "-Arg z1 \ 0" by auto - thus "acute_ang \\pi - arg z1\\ = acute_ang \\pi + arg z1\\" - using True canon_ang_plus_pi1[of "arg z1"] - using arg_bounded[of z1] canon_ang_plus_pi2[of "-arg z1"] + thus "acute_ang \\pi - Arg z1\\ = acute_ang \\pi + Arg z1\\" + using True canon_ang_plus_pi1[of "Arg z1"] + using Arg_bounded[of z1] canon_ang_plus_pi2[of "-Arg z1"] by (auto simp add:field_simps) next case False - hence "-arg z1 \ 0" + hence "-Arg z1 \ 0" by simp - thus "acute_ang \\pi - arg z1\\ = acute_ang \\pi + arg z1\\" - proof(cases "arg z1 = 0") + thus "acute_ang \\pi - Arg z1\\ = acute_ang \\pi + Arg z1\\" + proof(cases "Arg z1 = 0") case True thus ?thesis by (auto simp del:acute_ang_def) next case False - hence "-arg z1 > 0" - using \-arg z1 \ 0\ + hence "-Arg z1 > 0" + using \-Arg z1 \ 0\ by auto thus ?thesis - using False canon_ang_plus_pi1[of "-arg z1"] - using arg_bounded[of z1] canon_ang_plus_pi2[of "arg z1"] + using False canon_ang_plus_pi1[of "-Arg z1"] + using Arg_bounded[of z1] canon_ang_plus_pi2[of "Arg z1"] by (auto simp add:field_simps) qed qed next case True thus ?thesis using arg_cnj_pi[of z1] apply (auto simp del:acute_ang_def) - proof(cases "arg z2 = pi") + proof(cases "Arg z2 = pi") case True - thus "acute_ang \\arg z2 - pi\\ = acute_ang \\arg (cnj z2) - pi\\" + thus "acute_ang \\Arg z2 - pi\\ = acute_ang \\Arg (cnj z2) - pi\\" using arg_cnj_pi[of z2] by auto next case False - thus "acute_ang \\arg z2 - pi\\ = acute_ang \\arg (cnj z2) - pi\\" + thus "acute_ang \\Arg z2 - pi\\ = acute_ang \\Arg (cnj z2) - pi\\" using arg_cnj_not_pi[of z2] apply (auto simp del:acute_ang_def) - proof(cases "arg z2 > 0") + proof(cases "Arg z2 > 0") case True - hence "-arg z2 \ 0" + hence "-Arg z2 \ 0" by auto - thus "acute_ang \\arg z2 - pi\\ = acute_ang \\- arg z2 - pi\\" - using True canon_ang_minus_pi1[of "arg z2"] - using arg_bounded[of z2] canon_ang_minus_pi2[of "-arg z2"] + thus "acute_ang \\Arg z2 - pi\\ = acute_ang \\- Arg z2 - pi\\" + using True canon_ang_minus_pi1[of "Arg z2"] + using Arg_bounded[of z2] canon_ang_minus_pi2[of "-Arg z2"] by (auto simp add: field_simps) next case False - hence "-arg z2 \ 0" + hence "-Arg z2 \ 0" by simp - thus "acute_ang \\arg z2 - pi\\ = acute_ang \\- arg z2 - pi\\" - proof(cases "arg z2 = 0") + thus "acute_ang \\Arg z2 - pi\\ = acute_ang \\- Arg z2 - pi\\" + proof(cases "Arg z2 = 0") case True thus ?thesis by (auto simp del:acute_ang_def) next case False - hence "-arg z2 > 0" - using \-arg z2 \ 0\ + hence "-Arg z2 > 0" + using \-Arg z2 \ 0\ by auto thus ?thesis - using False canon_ang_minus_pi1[of "-arg z2"] - using arg_bounded[of z2] canon_ang_minus_pi2[of "arg z2"] + using False canon_ang_minus_pi1[of "-Arg z2"] + using Arg_bounded[of z2] canon_ang_minus_pi2[of "Arg z2"] by (auto simp add:field_simps) qed qed qed qed qed text \Cosine and scalar product\ lemma ortho_a_scalprod0: assumes "z1 \ 0" and "z2 \ 0" shows "\a z1 z2 = pi/2 \ scalprod z1 z2 = 0" unfolding ang_vec_a_def using assms ortho_c_scalprod0[of z1 z2] by auto declare ang_vec_c_def[simp del] lemma cos_a_c: "cos (\a z1 z2) = abs (cos (\c z1 z2))" proof- have "0 \ \c z1 z2" "\c z1 z2 \ pi" using ang_vec_c_bounded[of z1 z2] by auto show ?thesis proof (cases "\c z1 z2 = pi/2") case True thus ?thesis unfolding ang_vec_a_def acute_ang_def by (smt cos_pi_half pi_def pi_half) next case False show ?thesis proof (cases "\c z1 z2 < pi / 2") case True thus ?thesis using \0 \ \c z1 z2\ using cos_gt_zero_pi[of "\c z1 z2"] unfolding ang_vec_a_def by simp next case False hence "\c z1 z2 > pi/2" using \\c z1 z2 \ pi/2\ by simp hence "cos (\c z1 z2) < 0" using \\c z1 z2 \ pi\ using cos_lt_zero_on_pi2_pi[of "\c z1 z2"] by simp thus ?thesis using \\c z1 z2 > pi/2\ unfolding ang_vec_a_def by simp qed qed 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) + 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] + 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] + 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: 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)" + 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 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/Elementary_Complex_Geometry.thy b/thys/Complex_Geometry/Elementary_Complex_Geometry.thy --- a/thys/Complex_Geometry/Elementary_Complex_Geometry.thy +++ b/thys/Complex_Geometry/Elementary_Complex_Geometry.thy @@ -1,519 +1,519 @@ (* ----------------------------------------------------------------- *) section \Elementary complex geometry\ (* ----------------------------------------------------------------- *) text \In this section equations and basic properties of the most fundamental objects and relations in geometry -- collinearity, lines, circles and circlines. These are defined by equations in $\mathbb{C}$ (not extended by an infinite point). Later these equations will be generalized to equations in the extended complex plane, over homogenous coordinates.\ theory Elementary_Complex_Geometry imports More_Complex Linear_Systems Angles begin (* ----------------------------------------------------------------- *) subsection \Collinear points\ (* ----------------------------------------------------------------- *) definition collinear :: "complex \ complex \ complex \ bool" where "collinear z1 z2 z3 \ z1 = z2 \ Im ((z3 - z1) / (z2 - z1)) = 0" lemma collinear_ex_real: shows "collinear z1 z2 z3 \ (\ k::real. z1 = z2 \ z3 - z1 = complex_of_real k * (z2 - z1))" unfolding collinear_def by (metis Im_complex_of_real add_diff_cancel_right' complex_eq diff_zero legacy_Complex_simps(15) nonzero_mult_div_cancel_right right_minus_eq times_divide_eq_left zero_complex.code) text \Collinearity characterization using determinants\ lemma collinear_det: assumes "\ collinear z1 z2 z3" shows "det2 (z3 - z1) (cnj (z3 - z1)) (z1 - z2) (cnj (z1 - z2)) \ 0" proof- from assms have "((z3 - z1) / (z2 - z1)) - cnj ((z3 - z1) / (z2 - z1)) \ 0" "z2 \ z1" unfolding collinear_def using Complex_Im_express_cnj[of "(z3 - z1) / (z2 - z1)"] by (auto simp add: Complex_eq) thus ?thesis by (auto simp add: field_simps) qed text \Properties of three collinear points\ lemma collinear_sym1: shows "collinear z1 z2 z3 \ collinear z1 z3 z2" unfolding collinear_def using div_reals[of "1" "(z3 - z1)/(z2 - z1)"] div_reals[of "1" "(z2 - z1)/(z3 - z1)"] by auto lemma collinear_sym2': assumes "collinear z1 z2 z3" shows "collinear z2 z1 z3" proof- obtain k where "z1 = z2 \ z3 - z1 = complex_of_real k * (z2 - z1)" using assms unfolding collinear_ex_real by auto thus ?thesis proof assume "z3 - z1 = complex_of_real k * (z2 - z1)" thus ?thesis unfolding collinear_ex_real by (rule_tac x="1-k" in exI) (auto simp add: field_simps) qed (simp add: collinear_def) qed lemma collinear_sym2: shows "collinear z1 z2 z3 \ collinear z2 z1 z3" using collinear_sym2'[of z1 z2 z3] collinear_sym2'[of z2 z1 z3] by auto text \Properties of four collinear points\ lemma collinear_trans1: assumes "collinear z0 z2 z1" and "collinear z0 z3 z1" and "z0 \ z1" shows "collinear z0 z2 z3" using assms unfolding collinear_ex_real by (cases "z0 = z2", auto) (rule_tac x="k/ka" in exI, case_tac "ka = 0", auto simp add: field_simps) (* ----------------------------------------------------------------- *) subsection \Euclidean line\ (* ----------------------------------------------------------------- *) text \Line is defined by using collinearity\ definition line :: "complex \ complex \ complex set" where "line z1 z2 = {z. collinear z1 z2 z}" lemma line_points_collinear: assumes "z1 \ line z z'" and "z2 \ line z z'" and "z3 \ line z z'" and "z \ z'" shows "collinear z1 z2 z3" using assms unfolding line_def by (smt collinear_sym1 collinear_sym2' collinear_trans1 mem_Collect_eq) text \Parametric equation of a line\ lemma line_param: shows "z1 + cor k * (z2 - z1) \ line z1 z2" unfolding line_def by (auto simp add: collinear_def) text \Equation of the line containing two different given points\ lemma line_equation: assumes "z1 \ z2" and "\ = rot90 (z2 - z1)" shows "line z1 z2 = {z. cnj \*z + \*cnj z - (cnj \ * z1 + \ * cnj z1) = 0}" proof- { fix z have "z \ line z1 z2 \ Im ((z - z1)/(z2 - z1)) = 0" using assms by (simp add: line_def collinear_def) also have "... \ (z - z1)/(z2 - z1) = cnj ((z - z1)/(z2 - z1))" using complex_diff_cnj[of "(z - z1)/(z2 - z1)"] by auto also have "... \ (z - z1)*(cnj z2 - cnj z1) = (cnj z - cnj z1)*(z2 - z1)" using assms(1) using \(z \ line z1 z2) = is_real ((z - z1) / (z2 - z1))\ calculation is_real_div by auto also have "... \ cnj(z2 - z1)*z - (z2 - z1)*cnj z - (cnj(z2 - z1)*z1 - (z2 - z1)*cnj z1) = 0" by (simp add: field_simps) also have "... \ cnj \ * z + \ * cnj z - (cnj \ * z1 + \ * cnj z1) = 0" apply (subst assms)+ apply (subst cnj_mix_minus)+ by simp finally have "z \ line z1 z2 \ cnj \ * z + \ * cnj z - (cnj \ * z1 + \ * cnj z1) = 0" . } thus ?thesis by auto qed (* -------------------------------------------------------------------------- *) subsection \Euclidean circle\ (* -------------------------------------------------------------------------- *) text \Definition of the circle with given center and radius. It consists of all points on the distance $r$ from the center $\mu$.\ definition circle :: "complex \ real \ complex set" where "circle \ r = {z. cmod (z - \) = r}" text \Equation of the circle centered at $\mu$ with the radius $r$.\ lemma circle_equation: assumes "r \ 0" shows "circle \ r = {z. z*cnj z - z*cnj \ - cnj z*\ + \*cnj \ - cor (r*r) = 0}" proof (safe) fix z assume "z \ circle \ r" hence "(z - \)*cnj (z - \) = complex_of_real (r*r)" unfolding circle_def using complex_mult_cnj_cmod[of "z - \"] by (auto simp add: power2_eq_square) thus "z * cnj z - z * cnj \ - cnj z * \ + \ * cnj \ - cor (r * r) = 0" by (auto simp add: field_simps) next fix z assume "z * cnj z - z * cnj \ - cnj z * \ + \ * cnj \ - cor (r * r) = 0" hence "(z - \)*cnj (z - \) = cor (r*r)" by (auto simp add: field_simps) thus "z \ circle \ r" using assms using complex_mult_cnj_cmod[of "z - \"] using power2_eq_imp_eq[of "cmod (z - \)" r] unfolding circle_def power2_eq_square[symmetric] complex_of_real_def by auto qed (* -------------------------------------------------------------------------- *) subsection \Circline\ (* -------------------------------------------------------------------------- *) text \A very important property of the extended complex plane is that it is possible to treat circles and lines in a uniform way. The basic object is \emph{generalized circle}, or \emph{circline} for short. We introduce circline equation given in $\mathbb{C}$, and it will later be generalized to an equation in the extended complex plane $\overline{\mathbb{C}}$ given in matrix form using a Hermitean matrix and a quadratic form over homogenous coordinates.\ definition circline where "circline A BC D = {z. cor A*z*cnj z + cnj BC*z + BC*cnj z + cor D = 0}" text \Connection between circline and Euclidean circle\ text \Every circline with positive determinant and $A \neq 0$ represents an Euclidean circle\ lemma circline_circle: assumes "A \ 0" and "A * D \ (cmod BC)\<^sup>2" "cl = circline A BC D" and "\ = -BC/cor A" and "r2 = ((cmod BC)\<^sup>2 - A*D) / A\<^sup>2" and "r = sqrt r2" shows "cl = circle \ r" proof- have *: "cl = {z. z * cnj z + cnj (BC / cor A) * z + (BC / cor A) * cnj z + cor (D / A) = 0}" using \cl = circline A BC D\ \A \ 0\ by (auto simp add: circline_def field_simps) have "r2 \ 0" proof- have "(cmod BC)\<^sup>2 - A * D \ 0" using \A * D \ (cmod BC)\<^sup>2\ by auto thus ?thesis using \A \ 0\ \r2 = ((cmod BC)\<^sup>2 - A*D) / A\<^sup>2\ by (metis zero_le_divide_iff zero_le_power2) qed hence **: "r * r = r2" "r \ 0" using \r = sqrt r2\ by (auto simp add: real_sqrt_mult[symmetric]) have ***: "- \ * - cnj \ - cor r2 = cor (D / A)" using \\ = - BC / complex_of_real A\ \r2 = ((cmod BC)\<^sup>2 - A * D) / A\<^sup>2\ by (auto simp add: power2_eq_square complex_mult_cnj_cmod field_simps) (simp add: add_divide_eq_iff assms(1)) thus ?thesis using \r2 = ((cmod BC)\<^sup>2 - A*D) / A\<^sup>2\ \\ = - BC / cor A\ by (subst *, subst circle_equation[of r \, OF \r \ 0\], subst **) (auto simp add: field_simps power2_eq_square) qed lemma circline_ex_circle: assumes "A \ 0" and "A * D \ (cmod BC)\<^sup>2" and "cl = circline A BC D" shows "\ \ r. cl = circle \ r" using circline_circle[OF assms] by auto text \Every Euclidean circle can be represented by a circline\ lemma circle_circline: assumes "cl = circle \ r" and "r \ 0" shows "cl = circline 1 (-\) ((cmod \)\<^sup>2 - r\<^sup>2)" proof- have "complex_of_real ((cmod \)\<^sup>2 - r\<^sup>2) = \ * cnj \ - complex_of_real (r\<^sup>2)" by (auto simp add: complex_mult_cnj_cmod) thus "cl = circline 1 (- \) ((cmod \)\<^sup>2 - r\<^sup>2)" using assms using circle_equation[of r \] unfolding circline_def power2_eq_square by (simp add: field_simps) qed lemma circle_ex_circline: assumes "cl = circle \ r" and "r \ 0" shows "\ A BC D. A \ 0 \ A*D \ (cmod BC)\<^sup>2 \ cl = circline A BC D" using circle_circline[OF assms] using \r \ 0\ by (rule_tac x=1 in exI, rule_tac x="-\" in exI, rule_tac x="Re (\ * cnj \) - (r * r)" in exI) (simp add: complex_mult_cnj_cmod power2_eq_square) text \Connection between circline and Euclidean line\ text \Every circline with a positive determinant and $A = 0$ represents an Euclidean line\ lemma circline_line: assumes "A = 0" and "BC \ 0" and "cl = circline A BC D" and "z1 = - cor D * BC / (2 * BC * cnj BC)" and - "z2 = z1 + \ * sgn (if arg BC > 0 then -BC else BC)" + "z2 = z1 + \ * sgn (if Arg BC > 0 then -BC else BC)" shows "cl = line z1 z2" proof- have "cl = {z. cnj BC*z + BC*cnj z + complex_of_real D = 0}" using assms by (simp add: circline_def) have "{z. cnj BC*z + BC*cnj z + complex_of_real D = 0} = {z. cnj BC*z + BC*cnj z - (cnj BC*z1 + BC*cnj z1) = 0}" using \BC \ 0\ assms by simp moreover have "z1 \ z2" using \BC \ 0\ assms by (auto simp add: sgn_eq) moreover have "\ k. k \ 0 \ BC = cor k*rot90 (z2 - z1)" - proof (cases "arg BC > 0") + proof (cases "Arg BC > 0") case True thus ?thesis using assms by (rule_tac x="(cmod BC)" in exI, auto simp add: Complex_scale4) next case False thus ?thesis using assms by (rule_tac x="-(cmod BC)" in exI, simp) - (smt Complex.Re_sgn Im_sgn cis_arg complex_minus complex_surj mult_minus_right rcis_cmod_arg rcis_def) + (smt Complex.Re_sgn Im_sgn cis_Arg complex_minus complex_surj mult_minus_right rcis_cmod_Arg rcis_def) qed then obtain k where "cor k \ 0" "BC = cor k*rot90 (z2 - z1)" by auto moreover have *: "\ z. cnj_mix (BC / cor k) z - cnj_mix (BC / cor k) z1 = (1/cor k) * (cnj_mix BC z - cnj_mix BC z1)" using \cor k \ 0\ by (simp add: field_simps) hence "{z. cnj_mix BC z - cnj_mix BC z1 = 0} = {z. cnj_mix (BC / cor k) z - cnj_mix (BC / cor k) z1 = 0}" using \cor k \ 0\ by auto ultimately have "cl = line z1 z2" using line_equation[of z1 z2 "BC/cor k"] \cl = {z. cnj BC*z + BC*cnj z + complex_of_real D = 0}\ by auto thus ?thesis using \z1 \ z2\ by blast qed lemma circline_ex_line: assumes "A = 0" and "BC \ 0" and "cl = circline A BC D" shows "\ z1 z2. z1 \ z2 \ cl = line z1 z2" proof- let ?z1 = "- cor D * BC / (2 * BC * cnj BC)" - let ?z2 = "?z1 + \ * sgn (if 0 < arg BC then - BC else BC)" + let ?z2 = "?z1 + \ * sgn (if 0 < Arg BC then - BC else BC)" have "?z1 \ ?z2" using \BC \ 0\ by (simp add: sgn_eq) thus ?thesis using circline_line[OF assms, of ?z1 ?z2] \BC \ 0\ by (rule_tac x="?z1" in exI, rule_tac x="?z2" in exI, simp) qed text \Every Euclidean line can be represented by a circline\ lemma line_ex_circline: assumes "cl = line z1 z2" and "z1 \ z2" shows "\ BC D. BC \ 0 \ cl = circline 0 BC D" proof- let ?BC = "rot90 (z2 - z1)" let ?D = "Re (- 2 * scalprod z1 ?BC)" show ?thesis proof (rule_tac x="?BC" in exI, rule_tac x="?D" in exI, rule conjI) show "?BC \ 0" using \z1 \ z2\ rot90_ii[of "z2 - z1"] by auto next have *: "complex_of_real (Re (- 2 * scalprod z1 (rot90 (z2 - z1)))) = - (cnj_mix z1 (rot90 (z2 - z1)))" using rot90_ii[of "z2 - z1"] by (cases z1, cases z2, simp add: Complex_eq field_simps) show "cl = circline 0 ?BC ?D" apply (subst assms, subst line_equation[of z1 z2 ?BC]) unfolding circline_def by (fact, simp, subst *, simp add: field_simps) qed qed lemma circline_line': assumes "z1 \ z2" shows "circline 0 (\ * (z2 - z1)) (Re (- cnj_mix (\ * (z2 - z1)) z1)) = line z1 z2" proof- let ?B = "\ * (z2 - z1)" let ?D = "Re (- cnj_mix ?B z1)" have "circline 0 ?B ?D = {z. cnj ?B*z + ?B*cnj z + complex_of_real ?D = 0}" using assms by (simp add: circline_def) moreover have "is_real (- cnj_mix (\ * (z2 - z1)) z1)" using cnj_mix_real[of ?B z1] by auto hence "{z. cnj ?B*z + ?B*cnj z + complex_of_real ?D = 0} = {z. cnj ?B*z + ?B*cnj z - (cnj ?B*z1 + ?B*cnj z1) = 0}" apply (subst complex_of_real_Re, simp) unfolding diff_conv_add_uminus by simp moreover have "line z1 z2 = {z. cnj_mix (\ * (z2 - z1)) z - cnj_mix (\ * (z2 - z1)) z1 = 0}" using line_equation[of z1 z2 ?B] assms unfolding rot90_ii by simp ultimately show ?thesis by simp qed (* ---------------------------------------------------------------------------- *) subsection \Angle between two circles\ (* ---------------------------------------------------------------------------- *) text \Given a center $\mu$ of an Euclidean circle and a point $E$ on it, we define the tangent vector in $E$ as the radius vector $\overrightarrow{\mu E}$, rotated by $\pi/2$, clockwise or counterclockwise, depending on the circle orientation. The Boolean @{term p} encodes the orientation of the circle, and the function @{term "sgn_bool p"} returns $1$ when @{term p} is true, and $-1$ when @{term p} is false.\ abbreviation sgn_bool where "sgn_bool p \ if p then 1 else -1" definition circ_tang_vec :: "complex \ complex \ bool \ complex" where "circ_tang_vec \ E p = sgn_bool p * \ * (E - \)" text \Tangent vector is orthogonal to the radius.\ lemma circ_tang_vec_ortho: shows "scalprod (E - \) (circ_tang_vec \ E p) = 0" unfolding circ_tang_vec_def Let_def by auto text \Changing the circle orientation gives the opposite tangent vector.\ lemma circ_tang_vec_opposite_orient: shows "circ_tang_vec \ E p = - circ_tang_vec \ E (\ p)" unfolding circ_tang_vec_def by auto text \Angle between two oriented circles at their common point $E$ is defined as the angle between tangent vectors at $E$. Again we define three different angle measures.\ text \The oriented angle between two circles at the point $E$. The first circle is centered at $\mu_1$ and its orientation is given by the Boolean $p_1$, while the second circle is centered at $\mu_2$ and its orientation is given by the Boolea $p_2$.\ definition ang_circ where "ang_circ E \1 \2 p1 p2 = \ (circ_tang_vec \1 E p1) (circ_tang_vec \2 E p2)" text \The unoriented angle between the two circles\ definition ang_circ_c where "ang_circ_c E \1 \2 p1 p2 = \c (circ_tang_vec \1 E p1) (circ_tang_vec \2 E p2)" text \The acute angle between the two circles\ definition ang_circ_a where "ang_circ_a E \1 \2 p1 p2 = \a (circ_tang_vec \1 E p1) (circ_tang_vec \2 E p2)" text \Explicit expression for oriented angle between two circles\ lemma ang_circ_simp: assumes "E \ \1" and "E \ \2" shows "ang_circ E \1 \2 p1 p2 = - \arg (E - \2) - arg (E - \1) + sgn_bool p1 * pi / 2 - sgn_bool p2 * pi / 2\" + \Arg (E - \2) - Arg (E - \1) + sgn_bool p1 * pi / 2 - sgn_bool p2 * pi / 2\" unfolding ang_circ_def ang_vec_def circ_tang_vec_def apply (rule canon_ang_eq) using assms using arg_mult_2kpi[of "sgn_bool p2*\" "E - \2"] using arg_mult_2kpi[of "sgn_bool p1*\" "E - \1"] apply auto apply (rule_tac x="x-xa" in exI, auto simp add: field_simps) apply (rule_tac x="-1+x-xa" in exI, auto simp add: field_simps) apply (rule_tac x="1+x-xa" in exI, auto simp add: field_simps) apply (rule_tac x="x-xa" in exI, auto simp add: field_simps) done text \Explicit expression for the cosine of angle between two circles\ lemma cos_ang_circ_simp: assumes "E \ \1" and "E \ \2" shows "cos (ang_circ E \1 \2 p1 p2) = - sgn_bool (p1 = p2) * cos (arg (E - \2) - arg (E - \1))" + sgn_bool (p1 = p2) * cos (Arg (E - \2) - Arg (E - \1))" using assms - using cos_periodic_pi2[of "arg (E - \2) - arg (E - \1)"] - using cos_minus_pi[of "arg (E - \2) - arg (E - \1)"] + using cos_periodic_pi2[of "Arg (E - \2) - Arg (E - \1)"] + using cos_minus_pi[of "Arg (E - \2) - Arg (E - \1)"] using ang_circ_simp[OF assms, of p1 p2] by auto text \Explicit expression for the unoriented angle between two circles\ lemma ang_circ_c_simp: assumes "E \ \1" and "E \ \2" shows "ang_circ_c E \1 \2 p1 p2 = - \\arg (E - \2) - arg (E - \1) + sgn_bool p1 * pi / 2 - sgn_bool p2 * pi / 2\\" + \\Arg (E - \2) - Arg (E - \1) + sgn_bool p1 * pi / 2 - sgn_bool p2 * pi / 2\\" unfolding ang_circ_c_def ang_vec_c_def using ang_circ_simp[OF assms] unfolding ang_circ_def by auto text \Explicit expression for the acute angle between two circles\ lemma ang_circ_a_simp: assumes "E \ \1" and "E \ \2" shows "ang_circ_a E \1 \2 p1 p2 = - acute_ang (abs (canon_ang (arg(E - \2) - arg(E - \1) + (sgn_bool p1) * pi/2 - (sgn_bool p2) * pi/2)))" + acute_ang (abs (canon_ang (Arg(E - \2) - Arg(E - \1) + (sgn_bool p1) * pi/2 - (sgn_bool p2) * pi/2)))" unfolding ang_circ_a_def ang_vec_a_def using ang_circ_c_simp[OF assms] unfolding ang_circ_c_def by auto text \Acute angle between two circles does not depend on the circle orientation.\ lemma ang_circ_a_pTrue: assumes "E \ \1" and "E \ \2" shows "ang_circ_a E \1 \2 p1 p2 = ang_circ_a E \1 \2 True True" proof (cases "p1") case True show ?thesis proof (cases "p2") case True show ?thesis using \p1\ \p2\ by simp next case False show ?thesis using \p1\ \\ p2\ unfolding ang_circ_a_def using circ_tang_vec_opposite_orient[of \2 E p2] using ang_vec_a_opposite2 by simp qed next case False show ?thesis proof (cases "p2") case True show ?thesis using \\ p1\ \p2\ unfolding ang_circ_a_def using circ_tang_vec_opposite_orient[of \1 E p1] using ang_vec_a_opposite1 by simp next case False show ?thesis using \\ p1\ \\ p2\ unfolding ang_circ_a_def using circ_tang_vec_opposite_orient[of \1 E p1] circ_tang_vec_opposite_orient[of \2 E p2] using ang_vec_a_opposite1 ang_vec_a_opposite2 by simp qed qed text \Definition of the acute angle between the two unoriented circles \ abbreviation ang_circ_a' where "ang_circ_a' E \1 \2 \ ang_circ_a E \1 \2 True True" text \A very simple expression for the acute angle between the two circles\ lemma ang_circ_a_simp1: assumes "E \ \1" and "E \ \2" shows "ang_circ_a E \1 \2 p1 p2 = \a (E - \1) (E - \2)" unfolding ang_vec_a_def ang_vec_c_def ang_vec_def by (subst ang_circ_a_pTrue[OF assms, of p1 p2], subst ang_circ_a_simp[OF assms, of True True]) (metis add_diff_cancel) lemma ang_circ_a'_simp: assumes "E \ \1" and "E \ \2" shows "ang_circ_a' E \1 \2 = \a (E - \1) (E - \2)" by (rule ang_circ_a_simp1[OF assms]) 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 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)" + 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))" + 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 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 + 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,1089 +1,1066 @@ (* -------------------------------------------------------------------------- *) 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) +declare cis_cnj[simp] 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 cmod_cis [simp]: assumes "a \ 0" - shows "cor (cmod a) * cis (arg a) = a" + shows "of_real (cmod a) * cis (Arg a) = a" using assms - by (metis rcis_cmod_arg rcis_def) + 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) + shows "cis (Arg a) * of_real (cmod a) = a" + by (metis assms cmod_cis mult.commute) 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 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]: +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: +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: +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: +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 (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_cor_divide [simp]: shows "cmod (z / cor k) = cmod z / \k\" 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 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" 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" + 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_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: +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: +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]: +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]: +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]: +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]: +lemma cis_add_cis_opposite [simp]: shows "cis \ + cis (-\) = 2 * cos \" 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\ +text \@{term Arg} of real numbers\ lemma is_real_arg1: - assumes "arg z = 0 \ arg z = pi" + 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"] + 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" + shows "Arg z = 0 \ Arg z = pi" proof (cases "z = 0") case False thus ?thesis - using arg_bounded[of z] - by (smt (verit, best) Im_sgn assms cis.simps(2) cis_arg div_0 sin_zero_pi_iff) -qed (auto simp add: arg_zero) + using Arg_bounded[of z] + 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" + shows "Arg (cor k) = 0" proof- - have "cos (arg (Complex k 0)) > 0" + 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)"] + using rcis_cmod_Arg[of "Complex k 0"] Re_rcis[of "cmod (Complex k 0)" "Arg (Complex k 0)"] 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" + 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)"] + 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 (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) + 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) + 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\ +text \@{term Arg} of imaginary numbers\ lemma is_imag_arg1: - assumes "arg z = pi/2 \ arg z = -pi/2" + 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"] + 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" - using arg_bounded assms cos_0_iff_canon cos_arg_i_mult_zero by presburger + shows "Arg z = pi/2 \ Arg z = -pi/2" + 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" + 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)"] + 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 + 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" + 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)"] + 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 (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] + 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) + 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" + 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" - 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 + shows "\Arg z\ = Arg z" + using canon_ang_id[of "Arg z"] Arg_bounded by simp lemma arg_cis: - shows "arg (cis \) = \\\" + shows "Arg (cis \) = \\\" using arg_unique canon_ang canon_ang_cos canon_ang_sin cis.ctr sgn_cis by presburger -text \Cosine and sine of @{term arg}\ +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) + 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) + 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)" - by (metis assms cis_arg cis_mult mult_eq_0_iff sgn_mult) + shows "cis (Arg (z1 * z2)) = cis (Arg z1 + Arg z2)" + 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" + shows "\ k::int. Arg (z1 * z2) = Arg z1 + Arg z2 + 2*k*pi" proof- - have "cis (arg (z1*z2)) = cis (arg z1 + arg z2)" + 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"] + 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\" + shows "Arg(z1 * z2) = \Arg z1 + Arg z2\" proof- - obtain k::int where "arg(z1 * z2) = arg z1 + arg z2 + 2*k*pi" + 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\" + 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" + shows "Arg (cor k * z) = Arg z" proof (cases "z = 0") case False thus ?thesis using arg_mult assms canon_ang_arg by force -qed (auto simp: arg_zero) +qed (auto simp: Arg_zero) lemma arg_mult_real_negative [simp]: assumes "k < 0" - shows "arg (cor k * z) = arg (-z)" + shows "Arg (cor k * z) = Arg (-z)" proof (cases "z = 0") case False thus ?thesis using assms by (metis arg_mult_real_positive minus_mult_commute neg_0_less_iff_less of_real_minus minus_minus) -qed (auto simp: arg_zero) +qed (auto simp: Arg_zero) lemma arg_div_real_positive [simp]: assumes "k > 0" - shows "arg (z / cor k) = arg z" + 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)" + 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" - by (metis (no_types, lifting) arg_cis assms canon_ang_arg cis_arg mult_eq_0_iff nonzero_mult_div_cancel_left sgn_divide) + assumes "Arg (z * z1) = Arg (z * z2)" + shows "Arg z1 = Arg z2" + 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" + assumes "Arg z = pi" + shows "Arg (cnj z) = pi" 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") + 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 "arg (cnj z) = arg z \ arg(cnj z) = -arg z" - using arg_bounded[of z] arg_bounded[of "cnj z"] + have "Arg (cnj z) = Arg z \ Arg(cnj z) = -Arg z" + using Arg_bounded[of z] Arg_bounded[of "cnj z"] by (smt (verit, best) arccos_cos arccos_cos2 cnj.sel(1) complex_cnj_zero_iff complex_mod_cnj cos_arg) moreover - have "arg (cnj z) \ arg z" - using sin_0_iff_canon[of "arg (cnj z)"] arg_bounded False assms - by (metis complex_mod_cnj eq_cnj_iff_real is_real_arg2 rcis_cmod_arg) + have "Arg (cnj z) \ Arg z" + using sin_0_iff_canon[of "Arg (cnj z)"] Arg_bounded False assms + 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" + 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_pi[of z] \arg z = pi\ + 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" + 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) + 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\" + shows "Arg (1 / z) = \- Arg z\" 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" + 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" + 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" + 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\" + shows "Arg(z1 / z2) = \Arg z1 - Arg z2\" proof- - obtain k::int where "arg(z1 / z2) = arg z1 - arg z2 + 2*k*pi" + 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)" + 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\" + 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") + 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") + proof (cases "Arg z > 0") case True thus ?thesis using assms - using arg_bounded[of z] - using canon_ang_plus_pi1[of "arg z"] + 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 \Arg z \ 0\ using assms - using arg_bounded[of z] - using canon_ang_plus_pi2[of "arg z"] + 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)" + 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" + 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] + 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\ + using \\ = Arg z\ by auto qed subsubsection \Complex square root\ definition - "ccsqrt z = rcis (sqrt (cmod z)) (arg z / 2)" + "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) + 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" + 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 rcis_cmod_Arg[of z] rcis_cmod_Arg[of s] using arg_mult[of s s] - using canon_ang(3)[of "2*arg 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\ + 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)" + 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)" + 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)" + 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)" + assume ***: "cis (Arg s - k * pi) = cis (Arg s)" hence "s = ccsqrt z" - using rcis_cmod_arg[of s] + 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)" + assume ***: "cis (Arg s - k * pi) = -cis (Arg s)" hence "s = - ccsqrt z" - using rcis_cmod_arg[of s] + 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"] + 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" + 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)" + 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 (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" + 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" + 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") + 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] + 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"] + 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"] + 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" + 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) + 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: +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/Oriented_Circlines.thy b/thys/Complex_Geometry/Oriented_Circlines.thy --- a/thys/Complex_Geometry/Oriented_Circlines.thy +++ b/thys/Complex_Geometry/Oriented_Circlines.thy @@ -1,1373 +1,1373 @@ (* -------------------------------------------------------------------------- *) section \Oriented circlines\ (* -------------------------------------------------------------------------- *) theory Oriented_Circlines imports Circlines begin (* ----------------------------------------------------------------- *) subsection \Oriented circlines definition\ (* ----------------------------------------------------------------- *) text \In this section we describe how the orientation is introduced for the circlines. Similarly as the set of circline points, the set of disc points is introduced using the quadratic form induced by the circline matrix --- the set of points of the circline disc is the set of points such that satisfy that $A\cdot z\cdot \overline{z} + B\cdot \overline{z} + C\cdot z + D < 0$, where $(A, B, C, D)$ is a circline matrix representative Hermitean matrix. As the set of disc points must be invariant to the choice of representative, it is clear that oriented circlines matrices are equivalent only if they are proportional by a positive real factor (recall that unoriented circline allowed arbitrary non-zero real factors).\ definition ocircline_eq_cmat :: "complex_mat \ complex_mat \ bool" where [simp]: "ocircline_eq_cmat A B \(\ k::real. k > 0 \ B = cor k *\<^sub>s\<^sub>m A)" lift_definition ocircline_eq_clmat :: "circline_mat \ circline_mat \ bool" is ocircline_eq_cmat done lemma ocircline_eq_cmat_id [simp]: shows "ocircline_eq_cmat H H" by (simp, rule_tac x=1 in exI, simp) quotient_type ocircline = circline_mat / ocircline_eq_clmat proof (rule equivpI) show "reflp ocircline_eq_clmat" unfolding reflp_def by transfer (auto, rule_tac x="1" in exI, simp) next show "symp ocircline_eq_clmat" unfolding symp_def by transfer (simp only: ocircline_eq_cmat_def, safe, rule_tac x="1/k" in exI, simp) next show "transp ocircline_eq_clmat" unfolding transp_def by transfer (simp only: ocircline_eq_cmat_def, safe, rule_tac x="k*ka" in exI, simp) qed (* ----------------------------------------------------------------- *) subsection \Points on oriented circlines\ (* ----------------------------------------------------------------- *) text \Boundary of the circline.\ lift_definition on_ocircline :: "ocircline \ complex_homo \ bool" is on_circline_clmat_hcoords by transfer (simp del: quad_form_def, (erule exE)+, simp add: quad_form_scale_m quad_form_scale_v del: quad_form_def) definition ocircline_set :: "ocircline \ complex_homo set" where "ocircline_set H = {z. on_ocircline H z}" lemma ocircline_set_I [simp]: assumes "on_ocircline H z" shows "z \ ocircline_set H" using assms unfolding ocircline_set_def by simp (* ----------------------------------------------------------------- *) subsection \Disc and disc complement - in and out points\ (* ----------------------------------------------------------------- *) text \Interior and the exterior of an oriented circline.\ definition in_ocircline_cmat_cvec :: "complex_mat \ complex_vec \ bool" where [simp]: "in_ocircline_cmat_cvec H z \ Re (quad_form z H) < 0" lift_definition in_ocircline_clmat_hcoords :: "circline_mat \ complex_homo_coords \ bool" is in_ocircline_cmat_cvec done lift_definition in_ocircline :: "ocircline \ complex_homo \ bool" is in_ocircline_clmat_hcoords proof transfer fix H H' z z' assume hh: "hermitean H \ H \ mat_zero" and "hermitean H' \ H' \ mat_zero" and "z \ vec_zero" and "z' \ vec_zero" assume "ocircline_eq_cmat H H'" and "z \\<^sub>v z'" then obtain k k' where *: "0 < k" "H' = cor k *\<^sub>s\<^sub>m H" "k' \ 0" "z' = k' *\<^sub>s\<^sub>v z" by auto hence "quad_form z' H' = cor k * cor ((cmod k')\<^sup>2) * quad_form z H" by (simp add: quad_form_scale_v quad_form_scale_m del: vec_cnj_sv quad_form_def) hence "Re (quad_form z' H') = k * (cmod k')\<^sup>2 * Re (quad_form z H)" using hh quad_form_hermitean_real[of H] by (simp add: power2_eq_square) thus "in_ocircline_cmat_cvec H z = in_ocircline_cmat_cvec H' z'" using \k > 0\ \k' \ 0\ using mult_less_0_iff by fastforce qed definition disc :: "ocircline \ complex_homo set" where "disc H = {z. in_ocircline H z}" lemma disc_I [simp]: assumes "in_ocircline H z" shows "z \ disc H" using assms unfolding disc_def by simp definition out_ocircline_cmat_cvec :: "complex_mat \ complex_vec \ bool" where [simp]: "out_ocircline_cmat_cvec H z \ Re (quad_form z H) > 0" lift_definition out_ocircline_clmat_hcoords :: "circline_mat \ complex_homo_coords \ bool" is out_ocircline_cmat_cvec done lift_definition out_ocircline :: "ocircline \ complex_homo \ bool" is out_ocircline_clmat_hcoords proof transfer fix H H' z z' assume hh: "hermitean H \ H \ mat_zero" "hermitean H' \ H' \ mat_zero" "z \ vec_zero" "z' \ vec_zero" assume "ocircline_eq_cmat H H'" "z \\<^sub>v z'" then obtain k k' where *: "0 < k" "H' = cor k *\<^sub>s\<^sub>m H" "k' \ 0" "z' = k' *\<^sub>s\<^sub>v z" by auto hence "quad_form z' H' = cor k * cor ((cmod k')\<^sup>2) * quad_form z H" by (simp add: quad_form_scale_v quad_form_scale_m del: vec_cnj_sv quad_form_def) hence "Re (quad_form z' H') = k * (cmod k')\<^sup>2 * Re (quad_form z H)" using hh quad_form_hermitean_real[of H] by (simp add: power2_eq_square) thus "out_ocircline_cmat_cvec H z = out_ocircline_cmat_cvec H' z'" using \k > 0\ \k' \ 0\ using zero_less_mult_pos by fastforce qed definition disc_compl :: "ocircline \ complex_homo set" where "disc_compl H = {z. out_ocircline H z}" text \These three sets are mutually disjoint and they fill up the entire plane.\ lemma disc_compl_I [simp]: assumes "out_ocircline H z" shows "z \ disc_compl H" using assms unfolding disc_compl_def by simp lemma in_on_out: shows "in_ocircline H z \ on_ocircline H z \ out_ocircline H z" apply (transfer, transfer) using quad_form_hermitean_real using complex_eq_if_Re_eq by auto lemma in_on_out_univ: shows "disc H \ disc_compl H \ ocircline_set H = UNIV" unfolding disc_def disc_compl_def ocircline_set_def using in_on_out[of H] by auto lemma disc_inter_disc_compl [simp]: shows "disc H \ disc_compl H = {}" unfolding disc_def disc_compl_def by auto (transfer, transfer, simp) lemma disc_inter_ocircline_set [simp]: shows "disc H \ ocircline_set H = {}" unfolding disc_def ocircline_set_def by auto (transfer, transfer, simp) lemma disc_compl_inter_ocircline_set [simp]: shows "disc_compl H \ ocircline_set H = {}" unfolding disc_compl_def ocircline_set_def by auto (transfer, transfer, simp) (* ----------------------------------------------------------------- *) subsection \Opposite orientation\ (* ----------------------------------------------------------------- *) text \Finding opposite circline is idempotent, and opposite circlines share the same set of points, but exchange disc and its complement.\ definition opposite_ocircline_cmat :: "complex_mat \ complex_mat" where [simp]: "opposite_ocircline_cmat H = (-1) *\<^sub>s\<^sub>m H" lift_definition opposite_ocircline_clmat :: "circline_mat \ circline_mat" is opposite_ocircline_cmat by (auto simp add: hermitean_def mat_adj_def mat_cnj_def) lift_definition opposite_ocircline :: "ocircline \ ocircline" is opposite_ocircline_clmat by transfer auto lemma opposite_ocircline_involution [simp]: shows "opposite_ocircline (opposite_ocircline H) = H" by (transfer, transfer) (auto, rule_tac x="1" in exI, simp) lemma on_circline_opposite_ocircline_cmat [simp]: assumes "hermitean H \ H \ mat_zero" and "z \ vec_zero" shows "on_circline_cmat_cvec (opposite_ocircline_cmat H) z = on_circline_cmat_cvec H z" using assms by (simp add: quad_form_scale_m del: quad_form_def) lemma on_circline_opposite_ocircline [simp]: shows "on_ocircline (opposite_ocircline H) z \ on_ocircline H z" using on_circline_opposite_ocircline_cmat by (transfer, transfer, simp) lemma ocircline_set_opposite_ocircline [simp]: shows "ocircline_set (opposite_ocircline H) = ocircline_set H" unfolding ocircline_set_def by auto lemma disc_compl_opposite_ocircline [simp]: shows "disc_compl (opposite_ocircline H) = disc H" unfolding disc_def disc_compl_def apply auto apply (transfer, transfer) apply (auto simp add: quad_form_scale_m simp del: quad_form_def) apply (transfer ,transfer) apply (auto simp add: quad_form_scale_m simp del: quad_form_def) done lemma disc_opposite_ocircline [simp]: shows "disc (opposite_ocircline H) = disc_compl H" using disc_compl_opposite_ocircline[of "opposite_ocircline H"] by simp (* ----------------------------------------------------------------- *) subsection \Positive orientation. Conversion between unoriented and oriented circlines\ (* ----------------------------------------------------------------- *) text \Given an oriented circline, one can trivially obtain its unoriented counterpart, and these two share the same set of points.\ lift_definition of_ocircline :: "ocircline \ circline" is "id::circline_mat \ circline_mat" by transfer (simp, erule exE, force) lemma of_ocircline_opposite_ocircline [simp]: shows "of_ocircline (opposite_ocircline H) = of_ocircline H" by (transfer, transfer) (simp, erule exE, rule_tac x="-1" in exI, simp) lemma on_ocircline_of_circline [simp]: shows "on_circline (of_ocircline H) z \ on_ocircline H z" by (transfer, transfer, simp) lemma circline_set_of_ocircline [simp]: shows "circline_set (of_ocircline H) = ocircline_set H" unfolding ocircline_set_def circline_set_def by (safe) (transfer, simp)+ lemma inj_of_ocircline: assumes "of_ocircline H = of_ocircline H'" shows "H = H' \ H = opposite_ocircline H'" using assms by (transfer, transfer) (simp, metis linorder_neqE_linordered_idom minus_of_real_eq_of_real_iff mult_minus1 mult_sm_distribution neg_0_equal_iff_equal neg_less_0_iff_less) lemma inj_ocircline_set: assumes "ocircline_set H = ocircline_set H'" and "ocircline_set H \ {}" shows "H = H' \ H = opposite_ocircline H'" proof- from assms have "circline_set (of_ocircline H) = circline_set (of_ocircline H')" "circline_set (of_ocircline H') \ {}" by auto hence "of_ocircline H = of_ocircline H'" by (simp add: inj_circline_set) thus ?thesis by (rule inj_of_ocircline) qed text \Positive orientation.\ text \Given a representative Hermitean matrix of a circline, it represents exactly one of the two possible oriented circlines. The choice of what should be called a positive orientation is arbitrary. We follow Schwerdtfeger \cite{schwerdtfeger}, use the leading coefficient $A$ as the first criterion, and say that circline matrices with $A > 0$ are called positively oriented, and with $A < 0$ negatively oriented. However, Schwerdtfeger did not discuss the possible case of $A = 0$ (the case of lines), so we had to extend his definition to achieve a total characterization.\ definition pos_oriented_cmat :: "complex_mat \ bool" where [simp]: "pos_oriented_cmat H \ (let (A, B, C, D) = H - in (Re A > 0 \ (Re A = 0 \ ((B \ 0 \ arg B > 0) \ (B = 0 \ Re D > 0)))))" + in (Re A > 0 \ (Re A = 0 \ ((B \ 0 \ Arg B > 0) \ (B = 0 \ Re D > 0)))))" lift_definition pos_oriented_clmat :: "circline_mat \ bool" is pos_oriented_cmat done lift_definition pos_oriented :: "ocircline \ bool" is pos_oriented_clmat by transfer (case_tac circline_mat1, case_tac circline_mat2, simp, erule exE, simp, metis mult_pos_pos zero_less_mult_pos) lemma pos_oriented: shows "pos_oriented H \ pos_oriented (opposite_ocircline 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 moreover hence "Re A = 0 \ Re D = 0 \ B \ 0" using hh hermitean_elems[of A B C D] by (cases A, cases D) (auto simp add: Complex_eq) moreover - have "B \ 0 \ \ 0 < arg B \ 0 < arg (- B)" - using canon_ang_plus_pi2[of "arg B"] arg_bounded[of B] + have "B \ 0 \ \ 0 < Arg B \ 0 < Arg (- B)" + using canon_ang_plus_pi2[of "Arg B"] Arg_bounded[of B] by (auto simp add: arg_uminus) ultimately show "pos_oriented_cmat H \ pos_oriented_cmat (opposite_ocircline_cmat H)" by auto qed lemma pos_oriented_opposite_ocircline_cmat [simp]: assumes "hermitean H \ H \ mat_zero" shows "pos_oriented_cmat (opposite_ocircline_cmat H) \ \ pos_oriented_cmat H" proof- obtain A B C D where HH: "H = (A, B, C, D)" by (cases H) auto moreover hence "Re A = 0 \ Re D = 0 \ B \ 0" using assms hermitean_elems[of A B C D] by (cases A, cases D) (auto simp add: Complex_eq) moreover - have "B \ 0 \ \ 0 < arg B \ 0 < arg (- B)" - using canon_ang_plus_pi2[of "arg B"] arg_bounded[of B] + have "B \ 0 \ \ 0 < Arg B \ 0 < Arg (- B)" + using canon_ang_plus_pi2[of "Arg B"] Arg_bounded[of B] by (auto simp add: arg_uminus) moreover - have "B \ 0 \ 0 < arg B \ \ 0 < arg (- B)" - using canon_ang_plus_pi1[of "arg B"] arg_bounded[of B] + have "B \ 0 \ 0 < Arg B \ \ 0 < Arg (- B)" + using canon_ang_plus_pi1[of "Arg B"] Arg_bounded[of B] by (auto simp add: arg_uminus) ultimately show "pos_oriented_cmat (opposite_ocircline_cmat H) = (\ pos_oriented_cmat H)" by simp (metis not_less_iff_gr_or_eq) qed lemma pos_oriented_opposite_ocircline [simp]: shows "pos_oriented (opposite_ocircline H) \ \ pos_oriented H" using pos_oriented_opposite_ocircline_cmat by (transfer, transfer, simp) lemma pos_oriented_circle_inf: assumes "\\<^sub>h \ ocircline_set H" shows "pos_oriented H \ \\<^sub>h \ disc H" using assms unfolding ocircline_set_def disc_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 "is_real A" using hh hermitean_elems by auto assume "\ on_circline_cmat_cvec H \\<^sub>v" thus "pos_oriented_cmat H = (\ in_ocircline_cmat_cvec H \\<^sub>v)" using HH \is_real A\ by (cases A) (auto simp add: vec_cnj_def Complex_eq) qed lemma pos_oriented_euclidean_circle: assumes "is_circle (of_ocircline H)" "(a, r) = euclidean_circle (of_ocircline H)" "circline_type (of_ocircline H) < 0" shows "pos_oriented H \ of_complex a \ disc H" using assms unfolding disc_def apply simp proof (transfer, transfer) fix H a r assume hh: "hermitean H \ H \ mat_zero" obtain A B C D where HH: "H = (A, B, C, D)" by (cases H) auto hence "is_real A" "is_real D" "C = cnj B" using hh hermitean_elems by auto assume *: "\ circline_A0_cmat (id H)" "(a, r) = euclidean_circle_cmat (id H)" "circline_type_cmat (id H) < 0" hence "A \ 0" "Re A \ 0" using HH \is_real A\ by (case_tac[!] A) (auto simp add: Complex_eq) have "Re (A*D - B*C) < 0" using \circline_type_cmat (id H) < 0\ HH by simp have **: "(A * (D * cnj A) - B * (C * cnj A)) / (A * cnj A) = (A*D - B*C) / A" using \A \ 0\ by (simp add: field_simps) hence ***: "0 < Re A \ Re ((A * (D * cnj A) - B * (C * cnj A)) / (A * cnj A)) < 0" using \is_real A\ \A \ 0\ \Re (A*D - B*C) < 0\ by (simp add: Re_divide_real divide_less_0_iff) have "Re D - Re (cnj B * B / cnj A) < Re ((C - cnj B * A / cnj A) * B / A)" if "Re A > 0" using HH * \is_real A\ that by simp (smt "**" "***" cnj.simps(1) cnj.simps(2) complex_eq diff_divide_distrib left_diff_distrib' minus_complex.simps(1) mult.commute nonzero_mult_div_cancel_right)? moreover have "Re A > 0" if "Re D - Re (cnj B * B / cnj A) < Re ((C - cnj B * A / cnj A) * B / A)" using HH * \is_real A\ that by simp (smt "**" "***" cnj.simps(1) cnj.simps(2) complex_eq diff_divide_distrib left_diff_distrib' minus_complex.simps(1) mult.commute nonzero_mult_div_cancel_right)? ultimately show "pos_oriented_cmat H = in_ocircline_cmat_cvec H (of_complex_cvec a)" using HH \Re A \ 0\ * \is_real A\ by (auto simp add: vec_cnj_def) qed text \Introduce positive orientation\ definition of_circline_cmat :: "complex_mat \ complex_mat" where [simp]: "of_circline_cmat H = (if pos_oriented_cmat H then H else opposite_ocircline_cmat H)" lift_definition of_circline_clmat :: "circline_mat \ circline_mat" is of_circline_cmat by (auto simp add: hermitean_def mat_adj_def mat_cnj_def) lemma of_circline_clmat_def': shows "of_circline_clmat H = (if pos_oriented_clmat H then H else opposite_ocircline_clmat H)" by transfer simp lemma pos_oriented_cmat_mult_positive': assumes "hermitean H1 \ H1 \ mat_zero" and "hermitean H2 \ H2 \ mat_zero" and "\k. k > 0 \ H2 = cor k *\<^sub>s\<^sub>m H1" and "pos_oriented_cmat H1" shows "pos_oriented_cmat H2" proof- obtain A1 B1 C1 D1 A2 B2 C2 D2 where HH: "H1 = (A1, B1, C1, D1)" "H2 = (A2, B2, C2, D2)" by (cases H1, cases H2) thus ?thesis using assms by fastforce qed lemma pos_oriented_cmat_mult_positive: assumes "hermitean H1 \ H1 \ mat_zero" and "hermitean H2 \ H2 \ mat_zero" and "\k. k > 0 \ H2 = cor k *\<^sub>s\<^sub>m H1" shows "pos_oriented_cmat H1 \ pos_oriented_cmat H2" proof- from assms(3) obtain k where "k > 0 \ H2 = cor k *\<^sub>s\<^sub>m H1" by auto hence "\k. k > 0 \ H1 = cor k *\<^sub>s\<^sub>m H2" by (rule_tac x="1/k" in exI, auto) thus ?thesis using assms pos_oriented_cmat_mult_positive' by blast qed lemma pos_oriented_cmat_mult_negative: assumes "hermitean H1 \ H1 \ mat_zero" and "hermitean H2 \ H2 \ mat_zero" and "\k. k < 0 \ H2 = cor k *\<^sub>s\<^sub>m H1" shows "pos_oriented_cmat H1 \ \ pos_oriented_cmat H2" using assms proof- obtain A B C D A1 B1 C1 D1 where *: "H1 = (A, B, C, D)" "H2 = (A1, B1, C1, D1)" by (cases H1, cases H2) auto hence **: "is_real A" "is_real D" "is_real A1" "is_real D1" "B = 0 \ C = 0" "B1 = 0 \ C1 = 0" using assms hermitean_elems[of A B C D] hermitean_elems[of A1 B1 C1 D1] by auto show ?thesis proof (rule iffI) assume H1: "pos_oriented_cmat H1" show "\ pos_oriented_cmat H2" proof (cases "Re A > 0") case True thus ?thesis using assms * ** mult_neg_pos by fastforce next case False show ?thesis proof (cases "B = 0") case True thus ?thesis using assms * ** H1 \\ Re A > 0\ mult_neg_pos by fastforce next case False thus ?thesis using arg_uminus_opposite_sign[of B] arg_mult_real_negative using assms * ** H1 \\ Re A > 0\ mult_neg_pos by fastforce qed qed next assume H2: "\ pos_oriented_cmat H2" show "pos_oriented_cmat H1" proof (cases "Re A > 0") case True thus ?thesis using * ** mult_neg_pos by fastforce next case False show ?thesis proof (cases "B = 0") case True thus ?thesis using assms * ** H2 \\ Re A > 0\ by simp (smt arg_0_iff arg_complex_of_real_negative arg_complex_of_real_positive arg_mult_eq complex_of_real_Re mult.right_neutral mult_eq_0_iff of_real_0 of_real_1 zero_complex.simps(1)) next case False thus ?thesis using assms \\ Re A > 0\ H2 * ** using arg_uminus_opposite_sign[of B] by (cases "Re A = 0", auto simp add: mult_neg_neg) qed qed qed qed lift_definition of_circline :: "circline \ ocircline" is of_circline_clmat proof transfer fix H1 H2 assume hh: "hermitean H1 \ H1 \ mat_zero" "hermitean H2 \ H2 \ mat_zero" assume "circline_eq_cmat H1 H2" then obtain k where *: "k \ 0 \ H2 = cor k *\<^sub>s\<^sub>m H1" by auto show "ocircline_eq_cmat (of_circline_cmat H1) (of_circline_cmat H2)" proof (cases "k > 0") case True hence "pos_oriented_cmat H1 = pos_oriented_cmat H2" using * pos_oriented_cmat_mult_positive[OF hh] by blast thus ?thesis using hh * \k > 0\ apply (simp del: pos_oriented_cmat_def) apply (rule conjI) apply (rule impI) apply (simp, rule_tac x=k in exI, simp) apply (rule impI) apply (simp, rule_tac x=k in exI, simp) done next case False hence "k < 0" using * by simp hence "pos_oriented_cmat H1 \ \ (pos_oriented_cmat H2)" using * pos_oriented_cmat_mult_negative[OF hh] by blast thus ?thesis using hh * \k < 0\ apply (simp del: pos_oriented_cmat_def) apply (rule conjI) apply (rule impI) apply (simp, rule_tac x="-k" in exI, simp) apply (rule impI) apply (simp, rule_tac x="-k" in exI, simp) done qed qed lemma pos_oriented_of_circline [simp]: shows "pos_oriented (of_circline H)" using pos_oriented_opposite_ocircline_cmat by (transfer, transfer, simp) lemma of_ocircline_of_circline [simp]: shows "of_ocircline (of_circline H) = H" apply (transfer, auto simp add: of_circline_clmat_def') apply (transfer, simp, rule_tac x="-1" in exI, simp) done lemma of_circline_of_ocircline_pos_oriented [simp]: assumes "pos_oriented H" shows "of_circline (of_ocircline H) = H" using assms by (transfer, transfer, simp, rule_tac x=1 in exI, simp) lemma inj_of_circline: assumes "of_circline H = of_circline H'" shows "H = H'" using assms proof (transfer, transfer) fix H H' assume "ocircline_eq_cmat (of_circline_cmat H) (of_circline_cmat H')" then obtain k where "k > 0" "of_circline_cmat H' = cor k *\<^sub>s\<^sub>m of_circline_cmat H" by auto thus "circline_eq_cmat H H'" using mult_sm_inv_l[of "-1" "H'" "cor k *\<^sub>s\<^sub>m H"] using mult_sm_inv_l[of "-1" "H'" "(- (cor k)) *\<^sub>s\<^sub>m H"] apply (simp split: if_split_asm) apply (rule_tac x="k" in exI, simp) apply (rule_tac x="-k" in exI, simp) apply (rule_tac x="-k" in exI, simp) apply (rule_tac x="k" in exI, simp) done qed lemma of_circline_of_ocircline: shows "of_circline (of_ocircline H') = H' \ of_circline (of_ocircline H') = opposite_ocircline H'" proof (cases "pos_oriented H'") case True thus ?thesis by auto next case False hence "pos_oriented (opposite_ocircline H')" using pos_oriented by auto thus ?thesis using of_ocircline_opposite_ocircline[of H'] using of_circline_of_ocircline_pos_oriented [of "opposite_ocircline H'"] by auto qed (* -------------------------------------------------------------------------- *) subsubsection \Set of points on oriented and unoriented circlines\ (* -------------------------------------------------------------------------- *) lemma ocircline_set_of_circline [simp]: shows "ocircline_set (of_circline H) = circline_set H" unfolding ocircline_set_def circline_set_def proof (safe) fix z assume "on_ocircline (of_circline H) z" thus "on_circline H z" by (transfer, transfer, simp del: on_circline_cmat_cvec_def opposite_ocircline_cmat_def split: if_split_asm) next fix z assume "on_circline H z" thus "on_ocircline (of_circline H) z" by (transfer, transfer, simp del: on_circline_cmat_cvec_def opposite_ocircline_cmat_def split: if_split_asm) qed (* ----------------------------------------------------------------- *) subsection \Some special oriented circlines and discs\ (* ----------------------------------------------------------------- *) lift_definition mk_ocircline :: "complex \ complex \ complex \ complex \ ocircline" is mk_circline_clmat done text \oriented unit circle and unit disc\ lift_definition ounit_circle :: "ocircline" is unit_circle_clmat done lemma pos_oriented_ounit_circle [simp]: shows "pos_oriented ounit_circle" by (transfer, transfer, simp) lemma of_ocircline_ounit_circle [simp]: shows "of_ocircline ounit_circle = unit_circle" by (transfer, transfer, simp) lemma of_circline_unit_circle [simp]: shows "of_circline (unit_circle) = ounit_circle" by (transfer, transfer, simp) lemma ocircline_set_ounit_circle [simp]: shows "ocircline_set ounit_circle = circline_set unit_circle" apply (subst of_circline_unit_circle[symmetric]) apply (subst ocircline_set_of_circline) apply simp done definition unit_disc :: "complex_homo set" where "unit_disc = disc ounit_circle" definition unit_disc_compl :: "complex_homo set" where "unit_disc_compl = disc_compl ounit_circle" definition unit_circle_set :: "complex_homo set" where "unit_circle_set = circline_set unit_circle" lemma zero_in_unit_disc [simp]: shows "0\<^sub>h \ unit_disc" unfolding unit_disc_def disc_def by (simp, transfer, transfer) (simp add: Let_def vec_cnj_def) lemma one_notin_unit_dic [simp]: shows "1\<^sub>h \ unit_disc" unfolding unit_disc_def disc_def by (simp, transfer, transfer) (simp add: Let_def vec_cnj_def) lemma inf_notin_unit_disc [simp]: shows "\\<^sub>h \ unit_disc" unfolding unit_disc_def disc_def by (simp, transfer, transfer) (simp add: Let_def vec_cnj_def) lemma unit_disc_iff_cmod_lt_1 [simp]: shows "of_complex c \ unit_disc \ cmod c < 1" unfolding unit_disc_def disc_def by (simp, transfer, transfer, simp add: vec_cnj_def cmod_def power2_eq_square) lemma unit_disc_cmod_square_lt_1 [simp]: assumes "z \ unit_disc" shows "(cmod (to_complex z))\<^sup>2 < 1" using assms inf_or_of_complex[of z] by (auto simp add: abs_square_less_1) lemma unit_disc_to_complex_inj: assumes "u \ unit_disc" and "v \ unit_disc" assumes "to_complex u = to_complex v" shows "u = v" using assms using inf_or_of_complex[of u] inf_or_of_complex[of v] by auto lemma inversion_unit_disc [simp]: shows "inversion ` unit_disc = unit_disc_compl" unfolding unit_disc_def unit_disc_compl_def disc_def disc_compl_def proof safe fix x assume "in_ocircline ounit_circle x" thus "out_ocircline ounit_circle (inversion x)" unfolding inversion_def by (transfer, transfer, auto simp add: vec_cnj_def) next fix x assume *: "out_ocircline ounit_circle x" show "x \ inversion ` Collect (in_ocircline ounit_circle)" proof (rule image_eqI) show "x = inversion (inversion x)" by auto next show "inversion x \ Collect (in_ocircline ounit_circle)" using * unfolding inversion_def by (simp, transfer, transfer, auto simp add: vec_cnj_def) qed qed lemma inversion_unit_disc_compl [simp]: shows "inversion ` unit_disc_compl = unit_disc" proof- have "inversion ` (inversion ` unit_disc) = unit_disc" by (auto simp del: inversion_unit_disc simp add: image_iff) thus ?thesis by simp qed lemma inversion_noteq_unit_disc: assumes "u \ unit_disc" and "v \ unit_disc" shows "inversion u \ v" proof- from assms have "inversion u \ unit_disc_compl" by (metis image_eqI inversion_unit_disc) thus ?thesis using assms unfolding unit_disc_def unit_disc_compl_def using disc_inter_disc_compl by fastforce qed lemma in_ocircline_ounit_circle_conjugate [simp]: assumes "in_ocircline ounit_circle z" shows "in_ocircline ounit_circle (conjugate z)" using assms by (transfer, transfer, auto simp add: vec_cnj_def) lemma conjugate_unit_disc [simp]: shows "conjugate ` unit_disc = unit_disc" unfolding unit_disc_def disc_def apply (auto simp add: image_iff) apply (rule_tac x="conjugate x" in exI, simp) done lemma conjugate_in_unit_disc [simp]: assumes "z \ unit_disc" shows "conjugate z \ unit_disc" using conjugate_unit_disc using assms by blast lemma out_ocircline_ounit_circle_conjugate [simp]: assumes "out_ocircline ounit_circle z" shows "out_ocircline ounit_circle (conjugate z)" using assms by (transfer, transfer, auto simp add: vec_cnj_def) lemma conjugate_unit_disc_compl [simp]: shows "conjugate ` unit_disc_compl = unit_disc_compl" unfolding unit_disc_compl_def disc_compl_def apply (auto simp add: image_iff) apply (rule_tac x="conjugate x" in exI, simp) done lemma conjugate_in_unit_disc_compl [simp]: assumes "z \ unit_disc_compl" shows "conjugate z \ unit_disc_compl" using conjugate_unit_disc_compl using assms by blast (* -------------------------------------------------------------------------- *) subsubsection \Oriented x axis and lower half plane\ (* -------------------------------------------------------------------------- *) lift_definition o_x_axis :: "ocircline" is x_axis_clmat done lemma o_x_axis_pos_oriented [simp]: shows "pos_oriented o_x_axis" by (transfer, transfer, simp) lemma of_ocircline_o_x_axis [simp]: shows "of_ocircline o_x_axis = x_axis" by (transfer, transfer, simp) lemma of_circline_x_axis [simp]: shows "of_circline x_axis = o_x_axis" using of_circline_of_ocircline_pos_oriented[of o_x_axis] using o_x_axis_pos_oriented by simp lemma ocircline_set_circline_set_x_axis [simp]: shows "ocircline_set o_x_axis = circline_set x_axis" by (subst of_circline_x_axis[symmetric], subst ocircline_set_of_circline, simp) lemma ii_in_disc_o_x_axis [simp]: shows "ii\<^sub>h \ disc o_x_axis" unfolding disc_def by simp (transfer, transfer, simp add: Let_def vec_cnj_def) lemma ii_notin_disc_o_x_axis [simp]: shows "ii\<^sub>h \ disc_compl o_x_axis" unfolding disc_compl_def by simp (transfer, transfer, simp add: Let_def vec_cnj_def) lemma of_complex_in_o_x_axis_disc [simp]: shows "of_complex z \ disc o_x_axis \ Im z < 0" unfolding disc_def by auto (transfer, transfer, simp add: vec_cnj_def)+ lemma inf_notin_disc_o_x_axis [simp]: shows "\\<^sub>h \ disc o_x_axis" unfolding disc_def by simp (transfer, transfer, simp add: vec_cnj_def) lemma disc_o_x_axis: shows "disc o_x_axis = of_complex ` {z. Im z < 0}" proof- { fix z assume "z \ disc o_x_axis" hence "\ x. Im x < 0 \ z = of_complex x" using inf_or_of_complex[of z] by auto } thus ?thesis by (auto simp add: image_iff) qed (* -------------------------------------------------------------------------- *) subsubsection \Oriented single point circline\ (* -------------------------------------------------------------------------- *) lift_definition o_circline_point_0 :: "ocircline" is circline_point_0_clmat done lemma of_ocircline_o_circline_point_0 [simp]: shows "of_ocircline o_circline_point_0 = circline_point_0" by (transfer, transfer, simp) (* ----------------------------------------------------------------- *) subsection \Möbius action on oriented circlines and discs\ (* ----------------------------------------------------------------- *) text \Möbius action on an oriented circline is the same as on to an unoriented circline.\ lift_definition moebius_ocircline :: "moebius \ ocircline \ ocircline" is moebius_circline_mmat_clmat apply (transfer, transfer) apply simp apply ((erule exE)+, (erule conjE)+) apply (simp add: mat_inv_mult_sm) apply (rule_tac x="ka / Re (k * cnj k)" in exI, auto simp add: complex_mult_cnj_cmod power2_eq_square) done text \Möbius action on (unoriented) circlines could have been defined using the action on oriented circlines, but not the other way around.\ lemma moebius_circline_ocircline: shows "moebius_circline M H = of_ocircline (moebius_ocircline M (of_circline H))" apply (transfer, simp add: of_circline_clmat_def', safe) apply (transfer, simp, rule_tac x="-1" in exI, simp) done lemma moebius_ocircline_circline: shows "moebius_ocircline M H = of_circline (moebius_circline M (of_ocircline H)) \ moebius_ocircline M H = opposite_ocircline (of_circline (moebius_circline M (of_ocircline H)))" apply (transfer, simp add: of_circline_clmat_def', safe) apply (transfer, simp, rule_tac x="1" in exI, simp) apply (transfer, simp, erule_tac x="1" in allE, simp) done text \Möbius action on oriented circlines have many nice properties as it was the case with Möbius action on (unoriented) circlines. These transformations are injective and form group under composition.\ lemma inj_moebius_ocircline [simp]: shows "inj (moebius_ocircline M)" unfolding inj_on_def proof (safe) fix H H' assume "moebius_ocircline M H = moebius_ocircline M H'" thus "H = H'" proof (transfer, transfer) fix M H H' :: complex_mat assume "mat_det M \ 0" let ?iM = "mat_inv M" assume "ocircline_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 simp del: congruence_def) thus "ocircline_eq_cmat H H'" using \mat_det M \ 0\ inj_congruence[of ?iM H' "cor k *\<^sub>s\<^sub>m H"] mat_det_inv[of M] by auto qed qed lemma moebius_ocircline_id_moebius [simp]: shows "moebius_ocircline id_moebius H = H" by (transfer, transfer) (force simp add: mat_adj_def mat_cnj_def) lemma moebius_ocircline_comp [simp]: shows "moebius_ocircline (moebius_comp M1 M2) H = moebius_ocircline M1 (moebius_ocircline M2 H)" by (transfer, transfer, simp, rule_tac x=1 in exI, simp add: mat_inv_mult_mm mult_mm_assoc) lemma moebius_ocircline_comp_inv_left [simp]: shows "moebius_ocircline (moebius_inv M) (moebius_ocircline M H) = H" by (subst moebius_ocircline_comp[symmetric]) simp lemma moebius_ocircline_comp_inv_right [simp]: shows "moebius_ocircline M (moebius_ocircline (moebius_inv M) H) = H" by (subst moebius_ocircline_comp[symmetric]) simp lemma moebius_ocircline_opposite_ocircline [simp]: shows "moebius_ocircline M (opposite_ocircline H) = opposite_ocircline (moebius_ocircline M H)" by (transfer, transfer, simp, rule_tac x=1 in exI, simp) text \Möbius action on oriented circlines preserve the set of points of the circline.\ lemma ocircline_set_moebius_ocircline [simp]: shows "ocircline_set (moebius_ocircline M H) = moebius_pt M ` ocircline_set H" (is "?lhs = ?rhs") proof- have "?rhs = circline_set (moebius_circline M (of_ocircline H))" by simp thus ?thesis using moebius_ocircline_circline[of M H] by auto qed lemma ocircline_set_fix_iff_ocircline_fix: assumes "ocircline_set H' \ {}" shows "ocircline_set (moebius_ocircline M H) = ocircline_set H' \ moebius_ocircline M H = H' \ moebius_ocircline M H = opposite_ocircline H'" using assms using inj_ocircline_set[of "moebius_ocircline M H" H'] by (auto simp del: ocircline_set_moebius_ocircline) lemma disc_moebius_ocircline [simp]: shows "disc (moebius_ocircline M H) = moebius_pt M ` (disc H)" proof (safe) fix z assume "z \ disc H" thus "moebius_pt M z \ disc (moebius_ocircline M H)" unfolding disc_def proof (safe) assume "in_ocircline H z" thus "in_ocircline (moebius_ocircline M H) (moebius_pt M z)" proof (transfer, transfer) fix H M :: complex_mat and z :: complex_vec assume "mat_det M \ 0" assume "in_ocircline_cmat_cvec H z" thus "in_ocircline_cmat_cvec (moebius_circline_cmat_cmat M H) (moebius_pt_cmat_cvec M z)" using \mat_det M \ 0\ quad_form_congruence[of M z] by simp qed qed next fix z assume "z \ disc (moebius_ocircline M H)" thus "z \ moebius_pt M ` disc H" unfolding disc_def proof(safe) assume "in_ocircline (moebius_ocircline M H) z" show "z \ moebius_pt M ` Collect (in_ocircline H)" proof show "z = moebius_pt M (moebius_pt (moebius_inv M) z)" by simp next show "moebius_pt (moebius_inv M) z \ Collect (in_ocircline H)" using \in_ocircline (moebius_ocircline M H) z\ proof (safe, transfer, transfer) fix M H :: complex_mat and z :: complex_vec assume "mat_det M \ 0" hence "congruence (mat_inv (mat_inv M)) (congruence (mat_inv M) H) = H" by (simp del: congruence_def) hence "quad_form z (congruence (mat_inv M) H) = quad_form (mat_inv M *\<^sub>m\<^sub>v z) H" using quad_form_congruence[of "mat_inv M" "z" "congruence (mat_inv M) H"] using \mat_det M \ 0\ mat_det_inv[of "M"] by simp moreover assume "in_ocircline_cmat_cvec (moebius_circline_cmat_cmat M H) z" ultimately show "in_ocircline_cmat_cvec H (moebius_pt_cmat_cvec (moebius_inv_cmat M) z)" by simp qed qed qed qed lemma disc_compl_moebius_ocircline [simp]: shows "disc_compl (moebius_ocircline M H) = moebius_pt M ` (disc_compl H)" proof (safe) fix z assume "z \ disc_compl H" thus "moebius_pt M z \ disc_compl (moebius_ocircline M H)" unfolding disc_compl_def proof (safe) assume "out_ocircline H z" thus "out_ocircline (moebius_ocircline M H) (moebius_pt M z)" proof (transfer, transfer) fix H M :: complex_mat and z :: complex_vec assume "mat_det M \ 0" assume "out_ocircline_cmat_cvec H z" thus "out_ocircline_cmat_cvec (moebius_circline_cmat_cmat M H) (moebius_pt_cmat_cvec M z)" using \mat_det M \ 0\ quad_form_congruence[of M z] by simp qed qed next fix z assume "z \ disc_compl (moebius_ocircline M H)" thus "z \ moebius_pt M ` disc_compl H" unfolding disc_compl_def proof(safe) assume "out_ocircline (moebius_ocircline M H) z" show "z \ moebius_pt M ` Collect (out_ocircline H)" proof show "z = moebius_pt M (moebius_pt (moebius_inv M) z)" by simp next show "moebius_pt (moebius_inv M) z \ Collect (out_ocircline H)" using \out_ocircline (moebius_ocircline M H) z\ proof (safe, transfer, transfer) fix M H :: complex_mat and z :: complex_vec assume "mat_det M \ 0" hence "congruence (mat_inv (mat_inv M)) (congruence (mat_inv M) H) = H" by (simp del: congruence_def) hence "quad_form z (congruence (mat_inv M) H) = quad_form (mat_inv M *\<^sub>m\<^sub>v z) H" using quad_form_congruence[of "mat_inv M" "z" "congruence (mat_inv M) H"] using \mat_det M \ 0\ mat_det_inv[of "M"] by simp moreover assume "out_ocircline_cmat_cvec (moebius_circline_cmat_cmat M H) z" ultimately show "out_ocircline_cmat_cvec H (moebius_pt_cmat_cvec (moebius_inv_cmat M) z)" by simp qed qed qed qed (* ----------------------------------------------------------------- *) subsection \Orientation after Möbius transformations\ (* ----------------------------------------------------------------- *) text \All Euclidean similarities preserve circline orientation.\ lemma moebius_similarity_oriented_lines_to_oriented_lines: assumes "a \ 0" shows "\\<^sub>h \ ocircline_set H \ \\<^sub>h \ ocircline_set (moebius_ocircline (moebius_similarity a b) H)" using moebius_similarity_lines_to_lines[OF \a \ 0\, of b "of_ocircline H"] by simp lemma moebius_similarity_preserve_orientation': assumes "a \ 0" and "\\<^sub>h \ ocircline_set H" and "pos_oriented H" shows "pos_oriented (moebius_ocircline (moebius_similarity a b) H)" proof- let ?M = "moebius_similarity a b" let ?H = "moebius_ocircline ?M H" have "\\<^sub>h \ ocircline_set ?H" using \\\<^sub>h \ ocircline_set H\ moebius_similarity_oriented_lines_to_oriented_lines[OF \a \ 0\] by simp have "\\<^sub>h \ disc_compl H" using \\\<^sub>h \ ocircline_set H\ \pos_oriented H\ pos_oriented_circle_inf[of H] in_on_out unfolding disc_def disc_compl_def ocircline_set_def by auto hence "\\<^sub>h \ disc_compl ?H" using moebius_similarity_inf[OF \a \ 0\, of b] by force thus "pos_oriented ?H" using pos_oriented_circle_inf[of ?H] disc_inter_disc_compl[of ?H] \\\<^sub>h \ ocircline_set ?H\ by auto qed lemma moebius_similarity_preserve_orientation: assumes "a \ 0" and "\\<^sub>h \ ocircline_set H" shows "pos_oriented H \ pos_oriented(moebius_ocircline (moebius_similarity a b) H)" proof- let ?M = "moebius_similarity a b" let ?H = "moebius_ocircline ?M H" have "\\<^sub>h \ ocircline_set ?H" using \\\<^sub>h \ ocircline_set H\ moebius_similarity_oriented_lines_to_oriented_lines[OF \a \ 0\] by simp have *: "H = moebius_ocircline (- moebius_similarity a b) ?H" by simp show ?thesis using \a \ 0\ using moebius_similarity_preserve_orientation' [OF \a \ 0\ \\\<^sub>h \ ocircline_set H\] using moebius_similarity_preserve_orientation'[OF _ \\\<^sub>h \ ocircline_set ?H\, of "1/a" "-b/a"] using moebius_similarity_inv[of a b, OF \a \ 0\] * by auto qed lemma reciprocal_preserve_orientation: assumes "0\<^sub>h \ disc_compl H" shows "pos_oriented (moebius_ocircline moebius_reciprocal H)" proof- have "\\<^sub>h \ disc_compl (moebius_ocircline moebius_reciprocal H)" using assms by force thus "pos_oriented (moebius_ocircline moebius_reciprocal H)" using pos_oriented_circle_inf[of "moebius_ocircline moebius_reciprocal H"] using disc_inter_disc_compl[of "moebius_ocircline moebius_reciprocal H"] using disc_compl_inter_ocircline_set[of "moebius_ocircline moebius_reciprocal H"] by auto qed lemma reciprocal_not_preserve_orientation: assumes "0\<^sub>h \ disc H" shows "\ pos_oriented (moebius_ocircline moebius_reciprocal H)" proof- let ?H = "moebius_ocircline moebius_reciprocal H" have "\\<^sub>h \ disc ?H" using assms by force thus "\ pos_oriented ?H" using pos_oriented_circle_inf[of ?H] disc_inter_ocircline_set[of ?H] by auto qed text \Orientation of the image of a given oriented circline $H$ under a given Möbius transformation $M$ depends on whether the pole of $M$ (the point that $M$ maps to $\infty_{hc}$) lies in the disc or in the disc complement of $H$ (if it is on the set of $H$, then it maps onto a line and we do not discuss the orientation).\ lemma pole_in_disc: assumes "M = mk_moebius a b c d" and "c \ 0" and "a*d - b*c \ 0" assumes "is_pole M z" "z \ disc H" shows "\ pos_oriented (moebius_ocircline M H)" 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 "0\<^sub>h = moebius_pt (moebius_translation (d/c)) z" using pole_mk_moebius[of a b c d z] assms by simp have "z \ ocircline_set H" using \z \ disc H\ disc_inter_ocircline_set[of H] by blast hence "0\<^sub>h \ ocircline_set (moebius_ocircline ?t2 H)" using \0\<^sub>h = moebius_pt ?t2 z\ using moebius_pt_neq_I[of z _ ?t2] by force hence *: "\\<^sub>h \ ocircline_set (moebius_ocircline (?r + ?t2) H)" using \0\<^sub>h = moebius_pt (moebius_translation (d / c)) z\ by (metis circline_set_moebius_circline circline_set_moebius_circline_iff circline_set_of_ocircline moebius_pt_comp moebius_reciprocal ocircline_set_moebius_ocircline plus_moebius_def reciprocal_zero) hence **: "\\<^sub>h \ ocircline_set (moebius_ocircline (?rd + ?r + ?t2) H)" using \a*d - b*c \ 0\ \c \ 0\ unfolding moebius_rotation_dilatation_def using moebius_similarity_oriented_lines_to_oriented_lines[of _ "moebius_ocircline (?r + ?t2) H"] by (metis divide_eq_0_iff divisors_zero moebius_ocircline_comp plus_moebius_def right_minus_eq) have "\ pos_oriented (moebius_ocircline (?r + ?t2) H)" using pole_mk_moebius[of a b c d z] assms using reciprocal_not_preserve_orientation by force hence "\ pos_oriented (moebius_ocircline (?rd + ?r + ?t2) H)" using * using \a*d - b*c \ 0\ \c \ 0\ using moebius_similarity_preserve_orientation[of _ "moebius_ocircline (?r + ?t2) H"] unfolding moebius_rotation_dilatation_def by simp hence "\ pos_oriented (moebius_ocircline (?t1 + ?rd + ?r + ?t2) H)" using ** using moebius_similarity_preserve_orientation[of _ "moebius_ocircline (?rd + ?r + ?t2) H"] unfolding moebius_translation_def by simp thus ?thesis using assms by simp (subst moebius_decomposition, simp_all) qed lemma pole_in_disc_compl: assumes "M = mk_moebius a b c d" and "c \ 0" and "a*d - b*c \ 0" assumes "is_pole M z" and "z \ disc_compl H" shows "pos_oriented (moebius_ocircline M H)" 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 "0\<^sub>h = moebius_pt (moebius_translation (d/c)) z" using pole_mk_moebius[of a b c d z] assms by simp have "z \ ocircline_set H" using \z \ disc_compl H\ disc_compl_inter_ocircline_set[of H] by blast hence "0\<^sub>h \ ocircline_set (moebius_ocircline ?t2 H)" using \0\<^sub>h = moebius_pt ?t2 z\ using moebius_pt_neq_I[of z _ ?t2] by force hence *: "\\<^sub>h \ ocircline_set (moebius_ocircline (?r + ?t2) H)" using \0\<^sub>h = moebius_pt (moebius_translation (d / c)) z\ by (metis circline_set_moebius_circline circline_set_moebius_circline_iff circline_set_of_ocircline moebius_pt_comp moebius_reciprocal ocircline_set_moebius_ocircline plus_moebius_def reciprocal_zero) hence **: "\\<^sub>h \ ocircline_set (moebius_ocircline (?rd + ?r + ?t2) H)" using \a*d - b*c \ 0\ \c \ 0\ unfolding moebius_rotation_dilatation_def using moebius_similarity_oriented_lines_to_oriented_lines[of _ "moebius_ocircline (?r + ?t2) H"] by (metis divide_eq_0_iff divisors_zero moebius_ocircline_comp plus_moebius_def right_minus_eq) have "pos_oriented (moebius_ocircline (?r + ?t2) H)" using pole_mk_moebius[of a b c d z] assms using reciprocal_preserve_orientation by force hence "pos_oriented (moebius_ocircline (?rd + ?r + ?t2) H)" using * using \a*d - b*c \ 0\ \c \ 0\ using moebius_similarity_preserve_orientation[of _ "moebius_ocircline (?r + ?t2) H"] unfolding moebius_rotation_dilatation_def by simp hence "pos_oriented (moebius_ocircline (?t1 + ?rd + ?r + ?t2) H)" using ** using moebius_similarity_preserve_orientation[of _ "moebius_ocircline (?rd + ?r + ?t2) H"] unfolding moebius_translation_def by simp thus ?thesis using assms by simp (subst moebius_decomposition, simp_all) qed (* ----------------------------------------------------------------- *) subsection \Oriented circlines uniqueness\ (* ----------------------------------------------------------------- *) lemma ocircline_01inf: assumes "0\<^sub>h \ ocircline_set H \ 1\<^sub>h \ ocircline_set H \ \\<^sub>h \ ocircline_set H" shows "H = o_x_axis \ H = opposite_ocircline o_x_axis" proof- have "0\<^sub>h \ circline_set (of_ocircline H) \ 1\<^sub>h \ circline_set (of_ocircline H) \ \\<^sub>h \ circline_set (of_ocircline H)" using assms by simp hence "of_ocircline H = x_axis" using unique_circline_01inf' by auto thus "H = o_x_axis \ H = opposite_ocircline o_x_axis" by (metis inj_of_ocircline of_ocircline_o_x_axis) qed lemma unique_ocircline_01inf: shows "\! H. 0\<^sub>h \ ocircline_set H \ 1\<^sub>h \ ocircline_set H \ \\<^sub>h \ ocircline_set H \ ii\<^sub>h \ disc H" proof show "0\<^sub>h \ ocircline_set o_x_axis \ 1\<^sub>h \ ocircline_set o_x_axis \ \\<^sub>h \ ocircline_set o_x_axis \ ii\<^sub>h \ disc o_x_axis" by simp next fix H assume "0\<^sub>h \ ocircline_set H \ 1\<^sub>h \ ocircline_set H \ \\<^sub>h \ ocircline_set H \ ii\<^sub>h \ disc H" hence "0\<^sub>h \ ocircline_set H \ 1\<^sub>h \ ocircline_set H \ \\<^sub>h \ ocircline_set H" "ii\<^sub>h \ disc H" by auto hence "H = o_x_axis \ H = opposite_ocircline o_x_axis" using ocircline_01inf by simp thus "H = o_x_axis" using \ii\<^sub>h \ disc H\ by auto qed lemma unique_ocircline_set: assumes "A \ B" and "A \ C" and "B \ C" shows "\! H. pos_oriented H \ (A \ ocircline_set H \ B \ ocircline_set H \ C \ ocircline_set H)" proof- 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[OF assms] by auto let ?iM = "moebius_pt (moebius_inv M)" have **: "?iM 0\<^sub>h = A" "?iM 1\<^sub>h = B" "?iM \\<^sub>h = C" using * by (auto simp add: moebius_pt_invert) let ?H = "moebius_ocircline (moebius_inv M) o_x_axis" have 1: "A \ ocircline_set ?H" "B \ ocircline_set ?H" "C \ ocircline_set ?H" using ** by auto have 2: "\ H'. A \ ocircline_set H' \ B \ ocircline_set H' \ C \ ocircline_set H' \ H' = ?H \ H' = opposite_ocircline ?H" proof- fix H' let ?H' = "ocircline_set H'" and ?H'' = "ocircline_set (moebius_ocircline M H')" assume "A \ ocircline_set H' \ B \ ocircline_set H' \ C \ ocircline_set H'" hence "moebius_pt M A \ ?H''" "moebius_pt M B \ ?H''" "moebius_pt M C \ ?H''" by auto hence "0\<^sub>h \ ?H''" "1\<^sub>h \ ?H''" "\\<^sub>h \ ?H''" using * by auto hence "moebius_ocircline M H' = o_x_axis \ moebius_ocircline M H' = opposite_ocircline o_x_axis" using ocircline_01inf by auto hence "o_x_axis = moebius_ocircline M H' \ o_x_axis = opposite_ocircline (moebius_ocircline M H')" by auto thus "H' = ?H \ H' = opposite_ocircline ?H" proof assume *: "o_x_axis = moebius_ocircline M H'" show "H' = moebius_ocircline (moebius_inv M) o_x_axis \ H' = opposite_ocircline (moebius_ocircline (moebius_inv M) o_x_axis)" by (rule disjI1) (subst *, simp) next assume *: "o_x_axis = opposite_ocircline (moebius_ocircline M H')" show "H' = moebius_ocircline (moebius_inv M) o_x_axis \ H' = opposite_ocircline (moebius_ocircline (moebius_inv M) o_x_axis)" by (rule disjI2) (subst *, simp) qed qed show ?thesis (is "\! x. ?P x") proof (cases "pos_oriented ?H") case True show ?thesis proof show "?P ?H" using 1 True by auto next fix H assume "?P H" thus "H = ?H" using 1 2[of H] True by auto qed next case False let ?OH = "opposite_ocircline ?H" show ?thesis proof show "?P ?OH" using 1 False by auto next fix H assume "?P H" thus "H = ?OH" using False 2[of H] by auto qed qed qed lemma ocircline_set_0h: assumes "ocircline_set H = {0\<^sub>h}" shows "H = o_circline_point_0 \ H = opposite_ocircline (o_circline_point_0)" proof- have "of_ocircline H = circline_point_0" using assms using unique_circline_type_zero_0' card_eq1_circline_type_zero[of "of_ocircline H"] by auto thus ?thesis by (metis inj_of_ocircline of_ocircline_o_circline_point_0) qed end diff --git a/thys/Complex_Geometry/Unit_Circle_Preserving_Moebius.thy b/thys/Complex_Geometry/Unit_Circle_Preserving_Moebius.thy --- a/thys/Complex_Geometry/Unit_Circle_Preserving_Moebius.thy +++ b/thys/Complex_Geometry/Unit_Circle_Preserving_Moebius.thy @@ -1,1202 +1,1202 @@ (* ---------------------------------------------------------------------------- *) section \Unit circle preserving Möbius transformations\ (* ---------------------------------------------------------------------------- *) text \In this section we shall examine Möbius transformations that map the unit circle onto itself. We shall say that they fix or preserve the unit circle (although, they do not need to fix each of its points).\ theory Unit_Circle_Preserving_Moebius imports Unitary11_Matrices Moebius Oriented_Circlines begin (* ---------------------------------------------------------------------------- *) subsection \Möbius transformations that fix the unit circle\ (* ---------------------------------------------------------------------------- *) text \We define Möbius transformations that preserve unit circle as transformations represented by generalized unitary matrices with the $1-1$ signature (elements of the gruop $GU_{1,1}(2, \mathbb{C})$, defined earlier in the theory Unitary11Matrices).\ lift_definition unit_circle_fix_mmat :: "moebius_mat \ bool" is unitary11_gen done lift_definition unit_circle_fix :: "moebius \ bool" is unit_circle_fix_mmat apply transfer apply (auto simp del: mult_sm.simps) apply (simp del: mult_sm.simps add: unitary11_gen_mult_sm) apply (simp del: mult_sm.simps add: unitary11_gen_div_sm) done text \Our algebraic characterisation (by matrices) is geometrically correct.\ lemma unit_circle_fix_iff: shows "unit_circle_fix M \ moebius_circline M unit_circle = unit_circle" (is "?rhs = ?lhs") proof assume ?lhs thus ?rhs proof (transfer, transfer) fix M :: complex_mat assume "mat_det M \ 0" assume "circline_eq_cmat (moebius_circline_cmat_cmat M unit_circle_cmat) unit_circle_cmat" then obtain k where "k \ 0" "(1, 0, 0, -1) = cor k *\<^sub>s\<^sub>m congruence (mat_inv M) (1, 0, 0, -1)" by auto hence "(1/cor k, 0, 0, -1/cor k) = congruence (mat_inv M) (1, 0, 0, -1)" using mult_sm_inv_l[of "cor k" "congruence (mat_inv M) (1, 0, 0, -1)" ] by simp hence "congruence M (1/cor k, 0, 0, -1/cor k) = (1, 0, 0, -1)" using \mat_det M \ 0\ mat_det_inv[of M] using congruence_inv[of "mat_inv M" "(1, 0, 0, -1)" "(1/cor k, 0, 0, -1/cor k)"] by simp hence "congruence M (1, 0, 0, -1) = cor k *\<^sub>s\<^sub>m (1, 0, 0, -1)" using congruence_scale_m[of "M" "1/cor k" "(1, 0, 0, -1)"] using mult_sm_inv_l[of "1/ cor k" "congruence M (1, 0, 0, -1)" "(1, 0, 0, -1)"] \k \ 0\ by simp thus "unitary11_gen M" using \k \ 0\ unfolding unitary11_gen_def by simp qed next assume ?rhs thus ?lhs proof (transfer, transfer) fix M :: complex_mat assume "mat_det M \ 0" assume "unitary11_gen M" hence "unitary11_gen (mat_inv M)" using \mat_det M \ 0\ using unitary11_gen_mat_inv by simp thus " circline_eq_cmat (moebius_circline_cmat_cmat M unit_circle_cmat) unit_circle_cmat" unfolding unitary11_gen_real by auto (rule_tac x="1/k" in exI, simp) qed qed lemma circline_set_fix_iff_circline_fix: assumes "circline_set H' \ {}" shows "circline_set (moebius_circline M H) = circline_set H' \ moebius_circline M H = H'" using assms by auto (rule inj_circline_set, auto) lemma unit_circle_fix_iff_unit_circle_set: shows "unit_circle_fix M \ moebius_pt M ` unit_circle_set = unit_circle_set" proof- have "circline_set unit_circle \ {}" using one_in_unit_circle_set by auto thus ?thesis using unit_circle_fix_iff[of M] circline_set_fix_iff_circline_fix[of unit_circle M unit_circle] by (simp add: unit_circle_set_def) qed text \Unit circle preserving Möbius transformations form a group. \ lemma unit_circle_fix_id_moebius [simp]: shows "unit_circle_fix id_moebius" by (transfer, transfer, simp add: unitary11_gen_def mat_adj_def mat_cnj_def) lemma unit_circle_fix_moebius_add [simp]: assumes "unit_circle_fix M1" and "unit_circle_fix M2" shows "unit_circle_fix (M1 + M2)" using assms unfolding unit_circle_fix_iff by auto lemma unit_circle_fix_moebius_comp [simp]: assumes "unit_circle_fix M1" and "unit_circle_fix M2" shows "unit_circle_fix (moebius_comp M1 M2)" using unit_circle_fix_moebius_add[OF assms] by simp lemma unit_circle_fix_moebius_uminus [simp]: assumes "unit_circle_fix M" shows "unit_circle_fix (-M)" using assms unfolding unit_circle_fix_iff by (metis moebius_circline_comp_inv_left uminus_moebius_def) lemma unit_circle_fix_moebius_inv [simp]: assumes "unit_circle_fix M" shows "unit_circle_fix (moebius_inv M)" using unit_circle_fix_moebius_uminus[OF assms] by simp text \Unit circle fixing transforms preserve inverse points.\ lemma unit_circle_fix_moebius_pt_inversion [simp]: assumes "unit_circle_fix M" shows "moebius_pt M (inversion z) = inversion (moebius_pt M z)" using assms using symmetry_principle[of z "inversion z" unit_circle M] using unit_circle_fix_iff[of M, symmetric] using circline_symmetric_inv_homo_disc[of z] using circline_symmetric_inv_homo_disc'[of "moebius_pt M z" "moebius_pt M (inversion z)"] by metis (* -------------------------------------------------------------------------- *) subsection \Möbius transformations that fix the imaginary unit circle\ (* -------------------------------------------------------------------------- *) text \Only for completeness we show that Möbius transformations that preserve the imaginary unit circle are exactly those characterised by generalized unitary matrices (with the (2, 0) signature).\ lemma imag_unit_circle_fixed_iff_unitary_gen: assumes "mat_det (A, B, C, D) \ 0" shows "moebius_circline (mk_moebius A B C D) imag_unit_circle = imag_unit_circle \ unitary_gen (A, B, C, D)" (is "?lhs = ?rhs") proof assume ?lhs thus ?rhs using assms proof (transfer, transfer) fix A B C D :: complex let ?M = "(A, B, C, D)" and ?E = "(1, 0, 0, 1)" assume "circline_eq_cmat (moebius_circline_cmat_cmat (mk_moebius_cmat A B C D) imag_unit_circle_cmat) imag_unit_circle_cmat" "mat_det ?M \ 0" then obtain k where "k \ 0" "?E = cor k *\<^sub>s\<^sub>m congruence (mat_inv ?M) ?E" by auto hence "unitary_gen (mat_inv ?M)" using mult_sm_inv_l[of "cor k" "congruence (mat_inv ?M) ?E" "?E"] unfolding unitary_gen_def by (metis congruence_def divide_eq_0_iff eye_def mat_eye_r of_real_eq_0_iff one_neq_zero) thus "unitary_gen ?M" using unitary_gen_inv[of "mat_inv ?M"] \mat_det ?M \ 0\ by (simp del: mat_inv.simps) qed next assume ?rhs thus ?lhs using assms proof (transfer, transfer) fix A B C D :: complex let ?M = "(A, B, C, D)" and ?E = "(1, 0, 0, 1)" assume "unitary_gen ?M" "mat_det ?M \ 0" hence "unitary_gen (mat_inv ?M)" using unitary_gen_inv[of ?M] by simp then obtain k where "k \ 0" "mat_adj (mat_inv ?M) *\<^sub>m\<^sub>m (mat_inv ?M) = cor k *\<^sub>s\<^sub>m eye" using unitary_gen_real[of "mat_inv ?M"] mat_det_inv[of ?M] by auto hence *: "?E = (1 / cor k) *\<^sub>s\<^sub>m (mat_adj (mat_inv ?M) *\<^sub>m\<^sub>m (mat_inv ?M))" using mult_sm_inv_l[of "cor k" eye "mat_adj (mat_inv ?M) *\<^sub>m\<^sub>m (mat_inv ?M)"] by simp have "\k. k \ 0 \ (1, 0, 0, 1) = cor k *\<^sub>s\<^sub>m (mat_adj (mat_inv (A, B, C, D)) *\<^sub>m\<^sub>m (1, 0, 0, 1) *\<^sub>m\<^sub>m mat_inv (A, B, C, D))" using \mat_det ?M \ 0\ \k \ 0\ by (metis "*" Im_complex_of_real Re_complex_of_real \mat_adj (mat_inv ?M) *\<^sub>m\<^sub>m mat_inv ?M = cor k *\<^sub>s\<^sub>m eye\ complex_of_real_Re eye_def mat_eye_l mult_mm_assoc mult_mm_sm mult_sm_eye_mm of_real_1 of_real_divide of_real_eq_1_iff zero_eq_1_divide_iff) thus "circline_eq_cmat (moebius_circline_cmat_cmat (mk_moebius_cmat A B C D) imag_unit_circle_cmat) imag_unit_circle_cmat" using \mat_det ?M \ 0\ \k \ 0\ by (simp del: mat_inv.simps) qed qed (* -------------------------------------------------------------------------- *) subsection \Möbius transformations that fix the oriented unit circle and the unit disc\ (* -------------------------------------------------------------------------- *) text \Möbius transformations that fix the unit circle either map the unit disc onto itself or exchange it with its exterior. The transformations that fix the unit disc can be recognized from their matrices -- they have the form as before, but additionally it must hold that $|a|^2 > |b|^2$.\ definition unit_disc_fix_cmat :: "complex_mat \ bool" where [simp]: "unit_disc_fix_cmat M \ (let (A, B, C, D) = M in unitary11_gen (A, B, C, D) \ (B = 0 \ Re ((A*D)/(B*C)) > 1))" lift_definition unit_disc_fix_mmat :: "moebius_mat \ bool" is unit_disc_fix_cmat done lift_definition unit_disc_fix :: "moebius \ bool" is unit_disc_fix_mmat proof transfer fix M M' :: complex_mat assume det: "mat_det M \ 0" "mat_det M' \ 0" assume "moebius_cmat_eq M M'" then obtain k where *: "k \ 0" "M' = k *\<^sub>s\<^sub>m M" by auto hence **: "unitary11_gen M \ unitary11_gen M'" using unitary11_gen_mult_sm[of k M] unitary11_gen_div_sm[of k M] by auto obtain A B C D where MM: "(A, B, C, D) = M" by (cases M) auto obtain A' B' C' D' where MM': "(A', B', C', D') = M'" by (cases M') auto show "unit_disc_fix_cmat M = unit_disc_fix_cmat M'" using * ** MM MM' by auto qed text \Transformations that fix the unit disc also fix the unit circle.\ lemma unit_disc_fix_unit_circle_fix [simp]: assumes "unit_disc_fix M" shows "unit_circle_fix M" using assms by (transfer, transfer, auto) text \Transformations that preserve the unit disc preserve the orientation of the unit circle.\ lemma unit_disc_fix_iff_ounit_circle: shows "unit_disc_fix M \ moebius_ocircline M ounit_circle = ounit_circle" (is "?rhs \ ?lhs") proof assume *: ?lhs have "moebius_circline M unit_circle = unit_circle" apply (subst moebius_circline_ocircline[of M unit_circle]) apply (subst of_circline_unit_circle) apply (subst *) by simp hence "unit_circle_fix M" by (simp add: unit_circle_fix_iff) thus ?rhs using * proof (transfer, transfer) fix M :: complex_mat assume "mat_det M \ 0" let ?H = "(1, 0, 0, -1)" obtain A B C D where MM: "(A, B, C, D) = M" by (cases M) auto assume "unitary11_gen M" "ocircline_eq_cmat (moebius_circline_cmat_cmat M unit_circle_cmat) unit_circle_cmat" then obtain k where "0 < k" "?H = cor k *\<^sub>s\<^sub>m congruence (mat_inv M) ?H" by auto hence "congruence M ?H = cor k *\<^sub>s\<^sub>m ?H" using congruence_inv[of "mat_inv M" "?H" "(1/cor k) *\<^sub>s\<^sub>m ?H"] \mat_det M \ 0\ using mult_sm_inv_l[of "cor k" "congruence (mat_inv M) ?H" "?H"] using mult_sm_inv_l[of "1/cor k" "congruence M ?H"] using congruence_scale_m[of M "1/cor k" "?H"] using \\B. \1 / cor k \ 0; (1 / cor k) *\<^sub>s\<^sub>m congruence M (1, 0, 0, - 1) = B\ \ congruence M (1, 0, 0, - 1) = (1 / (1 / cor k)) *\<^sub>s\<^sub>m B\ by (auto simp add: mat_det_inv) then obtain a b k' where "k' \ 0" "M = k' *\<^sub>s\<^sub>m (a, b, cnj b, cnj a)" "sgn (Re (mat_det (a, b, cnj b, cnj a))) = 1" using unitary11_sgn_det_orientation'[of M k] \k > 0\ by auto moreover have "mat_det (a, b, cnj b, cnj a) \ 0" using \sgn (Re (mat_det (a, b, cnj b, cnj a))) = 1\ by (smt sgn_0 zero_complex.simps(1)) ultimately show "unit_disc_fix_cmat M" using unitary11_sgn_det[of k' a b M A B C D] using MM[symmetric] \k > 0\ \unitary11_gen M\ by (simp add: sgn_1_pos split: if_split_asm) qed next assume ?rhs thus ?lhs proof (transfer, transfer) fix M :: complex_mat assume "mat_det M \ 0" obtain A B C D where MM: "(A, B, C, D) = M" by (cases M) auto assume "unit_disc_fix_cmat M" hence "unitary11_gen M" "B = 0 \ 1 < Re (A * D / (B * C))" using MM[symmetric] by auto have "sgn (if B = 0 then 1 else sgn (Re (A * D / (B * C)) - 1)) = 1" using \B = 0 \ 1 < Re (A * D / (B * C))\ by auto then obtain k' where "k' > 0" "congruence M (1, 0, 0, -1) = cor k' *\<^sub>s\<^sub>m (1, 0, 0, -1)" using unitary11_orientation[OF \unitary11_gen M\ MM[symmetric]] by (auto simp add: sgn_1_pos) thus "ocircline_eq_cmat (moebius_circline_cmat_cmat M unit_circle_cmat) unit_circle_cmat" using congruence_inv[of M "(1, 0, 0, -1)" "cor k' *\<^sub>s\<^sub>m (1, 0, 0, -1)"] \mat_det M \ 0\ using congruence_scale_m[of "mat_inv M" "cor k'" "(1, 0, 0, -1)"] by auto qed qed text \Our algebraic characterisation (by matrices) is geometrically correct.\ lemma unit_disc_fix_iff [simp]: assumes "unit_disc_fix M" shows "moebius_pt M ` unit_disc = unit_disc" using assms using unit_disc_fix_iff_ounit_circle[of M] unfolding unit_disc_def by (subst disc_moebius_ocircline[symmetric], simp) lemma unit_disc_fix_discI [simp]: assumes "unit_disc_fix M" and "u \ unit_disc" shows "moebius_pt M u \ unit_disc" using unit_disc_fix_iff assms by blast text \Unit disc preserving transformations form a group.\ lemma unit_disc_fix_id_moebius [simp]: shows "unit_disc_fix id_moebius" by (transfer, transfer, simp add: unitary11_gen_def mat_adj_def mat_cnj_def) lemma unit_disc_fix_moebius_add [simp]: assumes "unit_disc_fix M1" and "unit_disc_fix M2" shows "unit_disc_fix (M1 + M2)" using assms unfolding unit_disc_fix_iff_ounit_circle by auto lemma unit_disc_fix_moebius_comp [simp]: assumes "unit_disc_fix M1" and "unit_disc_fix M2" shows "unit_disc_fix (moebius_comp M1 M2)" using unit_disc_fix_moebius_add[OF assms] by simp lemma unit_disc_fix_moebius_uminus [simp]: assumes "unit_disc_fix M" shows "unit_disc_fix (-M)" using assms unfolding unit_disc_fix_iff_ounit_circle by (metis moebius_ocircline_comp_inv_left uminus_moebius_def) lemma unit_disc_fix_moebius_inv [simp]: assumes "unit_disc_fix M" shows "unit_disc_fix (moebius_inv M)" using unit_disc_fix_moebius_uminus[OF assms] by simp (* -------------------------------------------------------------------------- *) subsection \Rotations are unit disc preserving transformations\ (* -------------------------------------------------------------------------- *) lemma unit_disc_fix_rotation [simp]: shows "unit_disc_fix (moebius_rotation \)" unfolding moebius_rotation_def moebius_similarity_def by (transfer, transfer, simp add: unitary11_gen_def mat_adj_def mat_cnj_def cis_mult) lemma moebius_rotation_unit_circle_fix [simp]: shows "moebius_pt (moebius_rotation \) u \ unit_circle_set \ u \ unit_circle_set" using moebius_pt_moebius_inv_in_set unit_circle_fix_iff_unit_circle_set by fastforce lemma ex_rotation_mapping_u_to_positive_x_axis: assumes "u \ 0\<^sub>h" and "u \ \\<^sub>h" shows "\ \. moebius_pt (moebius_rotation \) u \ positive_x_axis" proof- from assms obtain c where *: "u = of_complex c" using inf_or_of_complex by blast - have "is_real (cis (- arg c) * c)" "Re (cis (-arg c) * c) > 0" + have "is_real (cis (- Arg c) * c)" "Re (cis (-Arg c) * c) > 0" using "*" assms is_real_rot_to_x_axis positive_rot_to_x_axis of_complex_zero_iff by blast+ thus ?thesis using * - by (rule_tac x="-arg c" in exI) (simp add: positive_x_axis_def circline_set_x_axis) + by (rule_tac x="-Arg c" in exI) (simp add: positive_x_axis_def circline_set_x_axis) qed lemma ex_rotation_mapping_u_to_positive_y_axis: assumes "u \ 0\<^sub>h" and "u \ \\<^sub>h" shows "\ \. moebius_pt (moebius_rotation \) u \ positive_y_axis" proof- from assms obtain c where *: "u = of_complex c" using inf_or_of_complex by blast - have "is_imag (cis (pi/2 - arg c) * c)" "Im (cis (pi/2 - arg c) * c) > 0" + have "is_imag (cis (pi/2 - Arg c) * c)" "Im (cis (pi/2 - Arg c) * c) > 0" using "*" assms is_real_rot_to_x_axis positive_rot_to_x_axis of_complex_zero_iff by - (simp, simp, simp add: field_simps) thus ?thesis using * - by (rule_tac x="pi/2-arg c" in exI) (simp add: positive_y_axis_def circline_set_y_axis) + by (rule_tac x="pi/2-Arg c" in exI) (simp add: positive_y_axis_def circline_set_y_axis) qed lemma wlog_rotation_to_positive_x_axis: assumes in_disc: "u \ unit_disc" and not_zero: "u \ 0\<^sub>h" assumes preserving: "\\ u. \u \ unit_disc; u \ 0\<^sub>h; P (moebius_pt (moebius_rotation \) u)\ \ P u" assumes x_axis: "\x. \is_real x; 0 < Re x; Re x < 1\ \ P (of_complex x)" shows "P u" proof- from in_disc obtain \ where *: "moebius_pt (moebius_rotation \) u \ positive_x_axis" using ex_rotation_mapping_u_to_positive_x_axis[of u] using inf_notin_unit_disc not_zero by blast let ?Mu = "moebius_pt (moebius_rotation \) u" have "P ?Mu" proof- let ?x = "to_complex ?Mu" have "?Mu \ unit_disc" "?Mu \ 0\<^sub>h" "?Mu \ \\<^sub>h" using \u \ unit_disc\ \u \ 0\<^sub>h\ by auto hence "is_real (to_complex ?Mu)" "0 < Re ?x" "Re ?x < 1" using * unfolding positive_x_axis_def circline_set_x_axis by (auto simp add: cmod_eq_Re) thus ?thesis using x_axis[of ?x] \?Mu \ \\<^sub>h\ by simp qed thus ?thesis using preserving[OF in_disc] not_zero by simp qed lemma wlog_rotation_to_positive_x_axis': assumes not_zero: "u \ 0\<^sub>h" and not_inf: "u \ \\<^sub>h" assumes preserving: "\\ u. \u \ 0\<^sub>h; u \ \\<^sub>h; P (moebius_pt (moebius_rotation \) u)\ \ P u" assumes x_axis: "\x. \is_real x; 0 < Re x\ \ P (of_complex x)" shows "P u" proof- from not_zero and not_inf obtain \ where *: "moebius_pt (moebius_rotation \) u \ positive_x_axis" using ex_rotation_mapping_u_to_positive_x_axis[of u] using inf_notin_unit_disc by blast let ?Mu = "moebius_pt (moebius_rotation \) u" have "P ?Mu" proof- let ?x = "to_complex ?Mu" have "?Mu \ 0\<^sub>h" "?Mu \ \\<^sub>h" using \u \ \\<^sub>h\ \u \ 0\<^sub>h\ by auto hence "is_real (to_complex ?Mu)" "0 < Re ?x" using * unfolding positive_x_axis_def circline_set_x_axis by (auto simp add: cmod_eq_Re) thus ?thesis using x_axis[of ?x] \?Mu \ \\<^sub>h\ by simp qed thus ?thesis using preserving[OF not_zero not_inf] by simp qed lemma wlog_rotation_to_positive_y_axis: assumes in_disc: "u \ unit_disc" and not_zero: "u \ 0\<^sub>h" assumes preserving: "\\ u. \u \ unit_disc; u \ 0\<^sub>h; P (moebius_pt (moebius_rotation \) u)\ \ P u" assumes y_axis: "\x. \is_imag x; 0 < Im x; Im x < 1\ \ P (of_complex x)" shows "P u" proof- from in_disc and not_zero obtain \ where *: "moebius_pt (moebius_rotation \) u \ positive_y_axis" using ex_rotation_mapping_u_to_positive_y_axis[of u] using inf_notin_unit_disc by blast let ?Mu = "moebius_pt (moebius_rotation \) u" have "P ?Mu" proof- let ?y = "to_complex ?Mu" have "?Mu \ unit_disc" "?Mu \ 0\<^sub>h" "?Mu \ \\<^sub>h" using \u \ unit_disc\ \u \ 0\<^sub>h\ by auto hence "is_imag (to_complex ?Mu)" "0 < Im ?y" "Im ?y < 1" using * unfolding positive_y_axis_def circline_set_y_axis by (auto simp add: cmod_eq_Im) thus ?thesis using y_axis[of ?y] \?Mu \ \\<^sub>h\ by simp qed thus ?thesis using preserving[OF in_disc not_zero] by simp qed (* ---------------------------------------------------------------------------- *) subsection \Blaschke factors are unit disc preserving transformations\ (* ---------------------------------------------------------------------------- *) text \For a given point $a$, Blaschke factor transformations are of the form $k \cdot \left(\begin{array}{cc}1 & -a\\ -\overline{a} & 1\end{array}\right)$. It is a disc preserving Möbius transformation that maps the point $a$ to zero (by the symmetry principle, it must map the inverse point of $a$ to infinity).\ definition blaschke_cmat :: "complex \ complex_mat" where [simp]: "blaschke_cmat a = (if cmod a \ 1 then (1, -a, -cnj a, 1) else eye)" lift_definition blaschke_mmat :: "complex \ moebius_mat" is blaschke_cmat by simp lift_definition blaschke :: "complex \ moebius" is blaschke_mmat done lemma blaschke_0_id [simp]: "blaschke 0 = id_moebius" by (transfer, transfer, simp) lemma blaschke_a_to_zero [simp]: assumes "cmod a \ 1" shows "moebius_pt (blaschke a) (of_complex a) = 0\<^sub>h" using assms by (transfer, transfer, simp) lemma blaschke_inv_a_inf [simp]: assumes "cmod a \ 1" shows "moebius_pt (blaschke a) (inversion (of_complex a)) = \\<^sub>h" using assms unfolding inversion_def by (transfer, transfer) (simp add: vec_cnj_def, rule_tac x="1/(1 - a*cnj a)" in exI, simp) lemma blaschke_inf [simp]: assumes "cmod a < 1" and "a \ 0" shows "moebius_pt (blaschke a) \\<^sub>h = of_complex (- 1 / cnj a)" using assms by (transfer, transfer, simp add: complex_mod_sqrt_Re_mult_cnj) lemma blaschke_0_minus_a [simp]: assumes "cmod a \ 1" shows "moebius_pt (blaschke a) 0\<^sub>h = ~\<^sub>h (of_complex a)" using assms by (transfer, transfer, simp) lemma blaschke_unit_circle_fix [simp]: assumes "cmod a \ 1" shows "unit_circle_fix (blaschke a)" using assms by (transfer, transfer) (simp add: unitary11_gen_def mat_adj_def mat_cnj_def) lemma blaschke_unit_disc_fix [simp]: assumes "cmod a < 1" shows "unit_disc_fix (blaschke a)" using assms proof (transfer, transfer) fix a assume *: "cmod a < 1" show "unit_disc_fix_cmat (blaschke_cmat a)" proof (cases "a = 0") case True thus ?thesis by (simp add: unitary11_gen_def mat_adj_def mat_cnj_def) next case False hence "Re (a * cnj a) < 1" using * by (metis complex_mod_sqrt_Re_mult_cnj real_sqrt_lt_1_iff) hence "1 / Re (a * cnj a) > 1" using False by (smt complex_div_gt_0 less_divide_eq_1_pos one_complex.simps(1) right_inverse_eq) hence "Re (1 / (a * cnj a)) > 1" by (simp add: complex_is_Real_iff) thus ?thesis by (simp add: unitary11_gen_def mat_adj_def mat_cnj_def) qed qed lemma blaschke_unit_circle_fix': assumes "cmod a \ 1" shows "moebius_circline (blaschke a) unit_circle = unit_circle" using assms using blaschke_unit_circle_fix unit_circle_fix_iff by simp lemma blaschke_ounit_circle_fix': assumes "cmod a < 1" shows "moebius_ocircline (blaschke a) ounit_circle = ounit_circle" proof- have "Re (a * cnj a) < 1" using assms by (metis complex_mod_sqrt_Re_mult_cnj real_sqrt_lt_1_iff) thus ?thesis using assms using blaschke_unit_disc_fix unit_disc_fix_iff_ounit_circle by simp qed lemma moebius_pt_blaschke [simp]: assumes "cmod a \ 1" and "z \ 1 / cnj a" shows "moebius_pt (blaschke a) (of_complex z) = of_complex ((z - a) / (1 - cnj a * z))" using assms proof (cases "a = 0") case True thus ?thesis by auto next case False thus ?thesis using assms apply (transfer, transfer) apply (simp add: complex_mod_sqrt_Re_mult_cnj) apply (rule_tac x="1 / (1 - cnj a * z)" in exI) apply (simp add: field_simps) done qed (* -------------------------------------------------------------------------- *) subsubsection \Blaschke factors for a real point $a$\ (* -------------------------------------------------------------------------- *) text \If the point $a$ is real, the Blaschke factor preserve x-axis and the upper and the lower halfplane.\ lemma blaschke_real_preserve_x_axis [simp]: assumes "is_real a" and "cmod a < 1" shows "moebius_pt (blaschke a) z \ circline_set x_axis \ z \ circline_set x_axis" proof (cases "a = 0") case True thus ?thesis by simp next case False have "cmod a \ 1" using assms by linarith let ?a = "of_complex a" let ?i = "inversion ?a" let ?M = "moebius_pt (blaschke a)" have *: "?M ?a = 0\<^sub>h" "?M ?i = \\<^sub>h" "?M 0\<^sub>h = of_complex (-a)" using \cmod a \ 1\ blaschke_a_to_zero[of a] blaschke_inv_a_inf[of a] blaschke_0_minus_a[of a] by auto let ?Mx = "moebius_circline (blaschke a) x_axis" have "?a \ circline_set x_axis" "?i \ circline_set x_axis" "0\<^sub>h \ circline_set x_axis" using \is_real a\ \a \ 0\ eq_cnj_iff_real[of a] by auto hence "0\<^sub>h \ circline_set ?Mx" "\\<^sub>h \ circline_set ?Mx" "of_complex (-a) \ circline_set ?Mx" using * apply - apply (force simp add: image_iff)+ apply (simp add: image_iff, rule_tac x="0\<^sub>h" in bexI, simp_all) done moreover have "0\<^sub>h \ circline_set x_axis" "\\<^sub>h \ circline_set x_axis" "of_complex (-a) \ circline_set x_axis" using \is_real a\ by auto moreover have "of_complex (-a) \ 0\<^sub>h" using \a \ 0\ by simp hence "0\<^sub>h \ of_complex (-a)" by metis hence "\!H. 0\<^sub>h \ circline_set H \ \\<^sub>h \ circline_set H \ of_complex (- a) \ circline_set H" using unique_circline_set[of "0\<^sub>h" "\\<^sub>h" "of_complex (-a)"] \a \ 0\ by simp ultimately have "moebius_circline (blaschke a) x_axis = x_axis" by auto thus ?thesis by (metis circline_set_moebius_circline_iff) qed lemma blaschke_real_preserve_sgn_Im [simp]: assumes "is_real a" and "cmod a < 1" and "z \ \\<^sub>h" and "z \ inversion (of_complex a)" shows "sgn (Im (to_complex (moebius_pt (blaschke a) z))) = sgn (Im (to_complex z))" proof (cases "a = 0") case True thus ?thesis by simp next case False obtain z' where z': "z = of_complex z'" using inf_or_of_complex[of z] \z \ \\<^sub>h\ by auto have "z' \ 1 / cnj a" using assms z' \a \ 0\ by (auto simp add: of_complex_inj) moreover have "a * cnj a \ 1" using \cmod a < 1\ by auto (simp add: complex_mod_sqrt_Re_mult_cnj) moreover have "sgn (Im ((z' - a) / (1 - a * z'))) = sgn (Im z')" proof- have "a * z' \ 1" using \is_real a\ \z' \ 1 / cnj a\ \a \ 0\ eq_cnj_iff_real[of a] by (simp add: field_simps) moreover have "Re (1 - a\<^sup>2) > 0" using \is_real a\ \cmod a < 1\ by (smt Re_power2 minus_complex.simps(1) norm_complex_def one_complex.simps(1) power2_less_0 real_sqrt_lt_1_iff) moreover have "Im ((z' - a) / (1 - a * z')) = Re (((1 - a\<^sup>2) * Im z') / (cmod (1 - a*z'))\<^sup>2)" proof- have "1 - a * cnj z' \ 0" using \z' \ 1 / cnj a\ by (metis Im_complex_div_eq_0 complex_cnj_zero_iff diff_eq_diff_eq diff_numeral_special(9) eq_divide_imp is_real_div mult_not_zero one_complex.simps(2) zero_neq_one) hence "Im ((z' - a) / (1 - a * z')) = Im (((z' - a) * (1 - a * cnj z')) / ((1 - a * z') * cnj (1 - a * z')))" using \is_real a\ eq_cnj_iff_real[of a] by simp also have "... = Im ((z' - a - a * z' * cnj z' + a\<^sup>2 * cnj z') / (cmod (1 - a*z'))\<^sup>2)" unfolding complex_mult_cnj_cmod by (simp add: power2_eq_square field_simps) finally show ?thesis using \is_real a\ by (simp add: field_simps) qed moreover have "0 < (1 - (Re a)\<^sup>2) * Im z' / (cmod (1 - a * z'))\<^sup>2 \ Im z' > 0" using \is_real a\ \0 < Re (1 - a\<^sup>2)\ by (smt Re_power_real divide_le_0_iff minus_complex.simps(1) not_sum_power2_lt_zero one_complex.simps(1) zero_less_mult_pos) ultimately show ?thesis unfolding sgn_real_def using \cmod a < 1\ \a * z' \ 1\ \is_real a\ by (auto simp add: cmod_eq_Re) qed ultimately show ?thesis using assms z' moebius_pt_blaschke[of a z'] \is_real a\ eq_cnj_iff_real[of a] by simp qed lemma blaschke_real_preserve_sgn_arg [simp]: assumes "is_real a" and "cmod a < 1" and "z \ circline_set x_axis" - shows "sgn (arg (to_complex (moebius_pt (blaschke a) z))) = sgn (arg (to_complex z))" + shows "sgn (Arg (to_complex (moebius_pt (blaschke a) z))) = sgn (Arg (to_complex z))" proof- have "z \ \\<^sub>h" using assms using special_points_on_x_axis''(3) by blast moreover have "z \ inversion (of_complex a)" using assms by (metis calculation circline_equation_x_axis circline_set_x_axis_I conjugate_of_complex inversion_of_complex inversion_sym is_real_inversion o_apply of_complex_zero reciprocal_zero to_complex_of_complex) ultimately show ?thesis using blaschke_real_preserve_sgn_Im[OF assms(1) assms(2), of z] by (smt arg_Im_sgn assms(3) circline_set_x_axis_I norm_sgn of_complex_to_complex) qed (* -------------------------------------------------------------------------- *) subsubsection \Inverse Blaschke transform\ (* -------------------------------------------------------------------------- *) definition inv_blaschke_cmat :: "complex \ complex_mat" where [simp]: "inv_blaschke_cmat a = (if cmod a \ 1 then (1, a, cnj a, 1) else eye)" lift_definition inv_blaschke_mmat :: "complex \ moebius_mat" is inv_blaschke_cmat by simp lift_definition inv_blaschke :: "complex \ moebius" is inv_blaschke_mmat done lemma inv_blaschke_neg [simp]: "inv_blaschke a = blaschke (-a)" by (transfer, transfer) simp lemma inv_blaschke: assumes "cmod a \ 1" shows "blaschke a + inv_blaschke a = 0" apply simp apply (transfer, transfer) by auto (rule_tac x="1/(1 - a*cnj a)" in exI, simp) lemma ex_unit_disc_fix_mapping_u_to_zero: assumes "u \ unit_disc" shows "\ M. unit_disc_fix M \ moebius_pt M u = 0\<^sub>h" proof- from assms obtain c where *: "u = of_complex c" by (metis inf_notin_unit_disc inf_or_of_complex) hence "cmod c < 1" using assms unit_disc_iff_cmod_lt_1 by simp thus ?thesis using * by (rule_tac x="blaschke c" in exI) (smt blaschke_a_to_zero blaschke_ounit_circle_fix' unit_disc_fix_iff_ounit_circle) qed lemma wlog_zero: assumes in_disc: "u \ unit_disc" assumes preserving: "\ a u. \u \ unit_disc; cmod a < 1; P (moebius_pt (blaschke a) u)\ \ P u" assumes zero: "P 0\<^sub>h" shows "P u" proof- have *: "moebius_pt (blaschke (to_complex u)) u = 0\<^sub>h" by (smt blaschke_a_to_zero in_disc inf_notin_unit_disc of_complex_to_complex unit_disc_iff_cmod_lt_1) thus ?thesis using preserving[of u "to_complex u"] in_disc zero using inf_or_of_complex[of u] by auto qed lemma wlog_real_zero: assumes in_disc: "u \ unit_disc" and real: "is_real (to_complex u)" assumes preserving: "\ a u. \u \ unit_disc; is_real a; cmod a < 1; P (moebius_pt (blaschke a) u)\ \ P u" assumes zero: "P 0\<^sub>h" shows "P u" proof- have *: "moebius_pt (blaschke (to_complex u)) u = 0\<^sub>h" by (smt blaschke_a_to_zero in_disc inf_notin_unit_disc of_complex_to_complex unit_disc_iff_cmod_lt_1) thus ?thesis using preserving[of u "to_complex u"] in_disc zero real using inf_or_of_complex[of u] by auto qed lemma unit_disc_fix_transitive: assumes in_disc: "u \ unit_disc" and "u' \ unit_disc" shows "\ M. unit_disc_fix M \ moebius_pt M u = u'" proof- have "\ u \ unit_disc. \ M. unit_disc_fix M \ moebius_pt M u = u'" (is "?P u'") proof (rule wlog_zero) show "u' \ unit_disc" by fact next show "?P 0\<^sub>h" by (simp add: ex_unit_disc_fix_mapping_u_to_zero) next fix a u assume "cmod a < 1" and *: "?P (moebius_pt (blaschke a) u)" show "?P u" proof fix u' assume "u' \ unit_disc" then obtain M' where "unit_disc_fix M'" "moebius_pt M' u' = moebius_pt (blaschke a) u" using * by auto thus "\M. unit_disc_fix M \ moebius_pt M u' = u" using \cmod a < 1\ blaschke_unit_disc_fix[of a] using unit_disc_fix_moebius_comp[of "- blaschke a" "M'"] using unit_disc_fix_moebius_inv[of "blaschke a"] by (rule_tac x="(- (blaschke a)) + M'" in exI, simp) qed qed thus ?thesis using assms by auto qed (* -------------------------------------------------------------------------- *) subsection \Decomposition of unit disc preserving Möbius transforms\ (* -------------------------------------------------------------------------- *) text \Each transformation preserving unit disc can be decomposed to a rotation around the origin and a Blaschke factors that maps a point within the unit disc to zero.\ lemma unit_disc_fix_decompose_blaschke_rotation: assumes "unit_disc_fix M" shows "\ k \. cmod k < 1 \ M = moebius_rotation \ + blaschke k" using assms unfolding moebius_rotation_def moebius_similarity_def proof (simp, transfer, transfer) fix M assume *: "mat_det M \ 0" "unit_disc_fix_cmat M" then obtain k a b :: complex 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 have "a \ 0" using * ** by auto then obtain a' k' \ where ***: "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)" using ** unitary11_gen_cis_blaschke[of k M a b] by auto blast have "a' = 0 \ 1 < 1 / (cmod a')\<^sup>2" using * *** complex_mult_cnj_cmod[of a'] by simp hence "cmod a' < 1" by (smt less_divide_eq_1_pos norm_zero one_less_power one_power2 pos2) thus "\k. cmod k < 1 \ (\\. moebius_cmat_eq M (moebius_comp_cmat (mk_moebius_cmat (cis \) 0 0 1) (blaschke_cmat k)))" using *** apply (rule_tac x=a' in exI) apply simp apply (rule_tac x=\ in exI) apply simp apply (rule_tac x="1/k'" in exI) by auto qed lemma wlog_unit_disc_fix: assumes "unit_disc_fix M" assumes b: "\ k. cmod k < 1 \ P (blaschke k)" assumes r: "\ \. P (moebius_rotation \)" assumes comp: "\M1 M2. \unit_disc_fix M1; P M1; unit_disc_fix M2; P M2\ \ P (M1 + M2)" shows "P M" using assms using unit_disc_fix_decompose_blaschke_rotation[OF assms(1)] using blaschke_unit_disc_fix by auto lemma ex_unit_disc_fix_to_zero_positive_x_axis: assumes "u \ unit_disc" and "v \ unit_disc" and "u \ v" shows "\ M. unit_disc_fix M \ moebius_pt M u = 0\<^sub>h \ moebius_pt M v \ positive_x_axis" proof- from assms obtain B where *: "unit_disc_fix B" "moebius_pt B u = 0\<^sub>h" using ex_unit_disc_fix_mapping_u_to_zero by blast let ?v = "moebius_pt B v" have "?v \ unit_disc" using \v \ unit_disc\ * by auto hence "?v \ \\<^sub>h" using inf_notin_unit_disc by auto have "?v \ 0\<^sub>h" using \u \ v\ * by (metis moebius_pt_invert) obtain R where "unit_disc_fix R" "moebius_pt R 0\<^sub>h = 0\<^sub>h" "moebius_pt R ?v \ positive_x_axis" using ex_rotation_mapping_u_to_positive_x_axis[of ?v] \?v \ 0\<^sub>h\ \?v \ \\<^sub>h\ using moebius_pt_rotation_inf_iff moebius_pt_moebius_rotation_zero unit_disc_fix_rotation by blast thus ?thesis using * moebius_comp[of R B, symmetric] using unit_disc_fix_moebius_comp by (rule_tac x="R + B" in exI) (simp add: comp_def) qed lemma wlog_x_axis: assumes in_disc: "u \ unit_disc" "v \ unit_disc" assumes preserved: "\ M u v. \unit_disc_fix M; u \ unit_disc; v \ unit_disc; P (moebius_pt M u) (moebius_pt M v)\ \ P u v" assumes axis: "\ x. \is_real x; 0 \ Re x; Re x < 1\ \ P 0\<^sub>h (of_complex x)" shows "P u v" proof (cases "u = v") case True have "P u u" (is "?Q u") proof (rule wlog_zero[where P="?Q"]) show "u \ unit_disc" by fact next show "?Q 0\<^sub>h" using axis[of 0] by simp next fix a u assume "u \ unit_disc" "cmod a < 1" "?Q (moebius_pt (blaschke a) u)" thus "?Q u" using preserved[of "blaschke a" u u] using blaschke_unit_disc_fix[of a] by simp qed thus ?thesis using True by simp next case False from in_disc obtain M where *: "unit_disc_fix M" "moebius_pt M u = 0\<^sub>h" "moebius_pt M v \ positive_x_axis" using ex_unit_disc_fix_to_zero_positive_x_axis False by auto then obtain x where **: "moebius_pt M v = of_complex x" "is_real x" unfolding positive_x_axis_def circline_set_x_axis by auto moreover have "of_complex x \ unit_disc" using \unit_disc_fix M\ \v \ unit_disc\ ** using unit_disc_fix_discI by fastforce hence "0 < Re x" "Re x < 1" using \moebius_pt M v \ positive_x_axis\ ** by (auto simp add: positive_x_axis_def cmod_eq_Re) ultimately have "P 0\<^sub>h (of_complex x)" using \is_real x\ axis by auto thus ?thesis using preserved[OF *(1) assms(1-2)] *(2) **(1) by simp qed lemma wlog_positive_x_axis: assumes in_disc: "u \ unit_disc" "v \ unit_disc" "u \ v" assumes preserved: "\ M u v. \unit_disc_fix M; u \ unit_disc; v \ unit_disc; u \ v; P (moebius_pt M u) (moebius_pt M v)\ \ P u v" assumes axis: "\ x. \is_real x; 0 < Re x; Re x < 1\ \ P 0\<^sub>h (of_complex x)" shows "P u v" proof- have "u \ v \ P u v" (is "?Q u v") proof (rule wlog_x_axis) show "u \ unit_disc" "v \ unit_disc" by fact+ next fix M u v assume "unit_disc_fix M" "u \ unit_disc" "v \ unit_disc" "?Q (moebius_pt M u) (moebius_pt M v)" thus "?Q u v" using preserved[of M u v] using moebius_pt_invert by blast next fix x assume "is_real x" "0 \ Re x" "Re x < 1" thus "?Q 0\<^sub>h (of_complex x)" using axis[of x] of_complex_zero_iff[of x] complex.expand[of x 0] by fastforce qed thus ?thesis using \u \ v\ by simp qed (* -------------------------------------------------------------------------- *) subsection \All functions that fix the unit disc\ (* -------------------------------------------------------------------------- *) text \It can be proved that continuous functions that fix the unit disc are either actions of Möbius transformations that fix the unit disc (homographies), or are compositions of actions of Möbius transformations that fix the unit disc and the conjugation (antihomographies). We postulate this as a definition, but it this characterisation could also be formally shown (we do not need this for our further applications).\ definition unit_disc_fix_f where "unit_disc_fix_f f \ (\ M. unit_disc_fix M \ (f = moebius_pt M \ f = moebius_pt M \ conjugate))" text \Unit disc fixing functions really fix unit disc.\ lemma unit_disc_fix_f_unit_disc: assumes "unit_disc_fix_f M" shows "M ` unit_disc = unit_disc" using assms unfolding unit_disc_fix_f_def using image_comp by force text \Actions of unit disc fixing Möbius transformations (unit disc fixing homographies) are unit disc fixing functions.\ lemma unit_disc_fix_f_moebius_pt [simp]: assumes "unit_disc_fix M" shows "unit_disc_fix_f (moebius_pt M)" using assms unfolding unit_disc_fix_f_def by auto text \Compositions of unit disc fixing Möbius transformations and conjugation (unit disc fixing antihomographies) are unit disc fixing functions.\ lemma unit_disc_fix_conjugate_moebius [simp]: assumes "unit_disc_fix M" shows "unit_disc_fix (conjugate_moebius M)" proof- have "\a aa ab b. \1 < Re (a * b / (aa * ab)); \ 1 < Re (cnj a * cnj b / (cnj aa * cnj ab))\ \ aa = 0" by (metis cnj.simps(1) complex_cnj_divide complex_cnj_mult) thus ?thesis using assms by (transfer, transfer) (auto simp add: mat_cnj_def unitary11_gen_def mat_adj_def field_simps) qed lemma unit_disc_fix_conjugate_comp_moebius [simp]: assumes "unit_disc_fix M" shows "unit_disc_fix_f (conjugate \ moebius_pt M)" using assms apply (subst conjugate_moebius) apply (simp add: unit_disc_fix_f_def) apply (rule_tac x="conjugate_moebius M" in exI, simp) done text \Uniti disc fixing functions form a group under function composition.\ lemma unit_disc_fix_f_comp [simp]: assumes "unit_disc_fix_f f1" and "unit_disc_fix_f f2" shows "unit_disc_fix_f (f1 \ f2)" using assms apply (subst (asm) unit_disc_fix_f_def) apply (subst (asm) unit_disc_fix_f_def) proof safe fix M M' assume "unit_disc_fix M" "unit_disc_fix M'" thus "unit_disc_fix_f (moebius_pt M \ moebius_pt M')" unfolding unit_disc_fix_f_def by (rule_tac x="M + M'" in exI) auto next fix M M' assume "unit_disc_fix M" "unit_disc_fix M'" thus "unit_disc_fix_f (moebius_pt M \ (moebius_pt M' \ conjugate))" unfolding unit_disc_fix_f_def by (subst comp_assoc[symmetric])+ (rule_tac x="M + M'" in exI, auto) next fix M M' assume "unit_disc_fix M" "unit_disc_fix M'" thus "unit_disc_fix_f ((moebius_pt M \ conjugate) \ moebius_pt M')" unfolding unit_disc_fix_f_def by (subst comp_assoc, subst conjugate_moebius, subst comp_assoc[symmetric])+ (rule_tac x="M + conjugate_moebius M'" in exI, auto) next fix M M' assume "unit_disc_fix M" "unit_disc_fix M'" thus "unit_disc_fix_f ((moebius_pt M \ conjugate) \ (moebius_pt M' \ conjugate))" apply (subst comp_assoc[symmetric], subst comp_assoc) apply (subst conjugate_moebius, subst comp_assoc, subst comp_assoc) apply (simp add: unit_disc_fix_f_def) apply (rule_tac x="M + conjugate_moebius M'" in exI, auto) done qed lemma unit_disc_fix_f_inv: assumes "unit_disc_fix_f M" shows "unit_disc_fix_f (inv M)" using assms apply (subst (asm) unit_disc_fix_f_def) proof safe fix M assume "unit_disc_fix M" have "inv (moebius_pt M) = moebius_pt (-M)" by (rule ext) (simp add: moebius_inv) thus "unit_disc_fix_f (inv (moebius_pt M))" using \unit_disc_fix M\ unfolding unit_disc_fix_f_def by (rule_tac x="-M" in exI, simp) next fix M assume "unit_disc_fix M" have "inv (moebius_pt M \ conjugate) = conjugate \ inv (moebius_pt M)" by (subst o_inv_distrib, simp_all) also have "... = conjugate \ (moebius_pt (-M))" using moebius_inv by auto also have "... = moebius_pt (conjugate_moebius (-M)) \ conjugate" by (simp add: conjugate_moebius) finally show "unit_disc_fix_f (inv (moebius_pt M \ conjugate))" using \unit_disc_fix M\ unfolding unit_disc_fix_f_def by (rule_tac x="conjugate_moebius (-M)" in exI, simp) qed (* -------------------------------------------------------------------------- *) subsubsection \Action of unit disc fixing functions on circlines\ (* -------------------------------------------------------------------------- *) definition unit_disc_fix_f_circline where "unit_disc_fix_f_circline f H = (if \ M. unit_disc_fix M \ f = moebius_pt M then moebius_circline (THE M. unit_disc_fix M \ f = moebius_pt M) H else if \ M. unit_disc_fix M \ f = moebius_pt M \ conjugate then (moebius_circline (THE M. unit_disc_fix M \ f = moebius_pt M \ conjugate) \ conjugate_circline) H else H)" lemma unique_moebius_pt_conjugate: assumes "moebius_pt M1 \ conjugate = moebius_pt M2 \ conjugate" shows "M1 = M2" proof- from assms have "moebius_pt M1 = moebius_pt M2" using conjugate_conjugate_comp rewriteL_comp_comp2 by fastforce thus ?thesis using unique_moebius_pt by auto qed lemma unit_disc_fix_f_circline_direct: assumes "unit_disc_fix M" and "f = moebius_pt M" shows "unit_disc_fix_f_circline f H = moebius_circline M H" proof- have "M = (THE M. unit_disc_fix M \ f = moebius_pt M)" using assms using theI_unique[of "\ M. unit_disc_fix M \ f = moebius_pt M" M] using unique_moebius_pt[of M] by auto thus ?thesis using assms unfolding unit_disc_fix_f_circline_def by auto qed lemma unit_disc_fix_f_circline_indirect: assumes "unit_disc_fix M" and "f = moebius_pt M \ conjugate" shows "unit_disc_fix_f_circline f H = ((moebius_circline M) \ conjugate_circline) H" proof- have "\ (\ M. unit_disc_fix M \ f = moebius_pt M)" using assms homography_antihomography_exclusive[of f] unfolding is_homography_def is_antihomography_def is_moebius_def by auto moreover have "M = (THE M. unit_disc_fix M \ f = moebius_pt M \ conjugate)" using assms using theI_unique[of "\ M. unit_disc_fix M \ f = moebius_pt M \ conjugate" M] using unique_moebius_pt_conjugate[of M] by auto ultimately show ?thesis using assms unfolding unit_disc_fix_f_circline_def by metis qed text \Disc automorphisms - it would be nice to show that there are no disc automorphisms other than unit disc fixing homographies and antihomographies, but this part of the theory is not yet developed.\ definition is_disc_aut where "is_disc_aut f \ bij_betw f unit_disc unit_disc" end \ No newline at end of file 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] + 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) + 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) +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 (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/Lambda_Free_KBOs/Lambda_Free_TKBO_Coefs.thy b/thys/Lambda_Free_KBOs/Lambda_Free_TKBO_Coefs.thy --- a/thys/Lambda_Free_KBOs/Lambda_Free_TKBO_Coefs.thy +++ b/thys/Lambda_Free_KBOs/Lambda_Free_TKBO_Coefs.thy @@ -1,1999 +1,1997 @@ (* Title: The Graceful Transfinite Knuth-Bendix Order with Subterm Coefficients for Lambda-Free Higher-Order Terms Author: Heiko Becker , 2016 Author: Jasmin Blanchette , 2016 Author: Uwe Waldmann , 2016 Author: Daniel Wand , 2016 Maintainer: Jasmin Blanchette *) section \The Graceful Transfinite Knuth--Bendix Order with Subterm Coefficients for Lambda-Free Higher-Order Terms\ theory Lambda_Free_TKBO_Coefs imports Lambda_Free_KBO_Util Nested_Multisets_Ordinals.Signed_Syntactic_Ordinal abbrevs "=p" = "=\<^sub>p" and ">p" = ">\<^sub>p" and "\p" = "\\<^sub>p" and ">t" = ">\<^sub>t" and "\t" = "\\<^sub>t" and "!h" = "\<^sub>h" begin text \ This theory defines the graceful transfinite Knuth--Bendix order (KBO) with subterm coefficients for \\\-free higher-order terms. The proof was developed by copying that of the standard KBO and generalizing it along two axes:\ subterm coefficients and ordinals. Both features complicate the definitions and proofs substantially. \ subsection \Setup\ -hide_const (open) Complex.arg - locale tkbo_coefs = kbo_std_basis _ _ arity_sym arity_var wt_sym for arity_sym :: "'s \ enat" and arity_var :: "'v \ enat" and wt_sym :: "'s \ hmultiset" + fixes coef_sym :: "'s \ nat \ hmultiset" assumes coef_sym_gt_0: "coef_sym f i > 0" begin abbreviation \\<^sub>h :: hmultiset where "\\<^sub>h \ of_nat \" abbreviation \\<^sub>h :: hmultiset where "\\<^sub>h \ of_nat \" abbreviation arity_sym\<^sub>h :: "'s \ hmultiset" where "arity_sym\<^sub>h f \ hmset_of_enat (arity_sym f)" abbreviation arity_var\<^sub>h :: "'v \ hmultiset" where "arity_var\<^sub>h f \ hmset_of_enat (arity_var f)" abbreviation arity_hd\<^sub>h :: "('s, 'v) hd \ hmultiset" where "arity_hd\<^sub>h f \ hmset_of_enat (arity_hd f)" abbreviation arity\<^sub>h :: "('s, 'v) tm \ hmultiset" where "arity\<^sub>h s \ hmset_of_enat (arity s)" lemma arity\<^sub>h_conv: "arity\<^sub>h s = arity_hd\<^sub>h (head s) - of_nat (num_args s)" unfolding arity_def by simp lemma arity\<^sub>h_App[simp]: "arity\<^sub>h (App s t) = arity\<^sub>h s - 1" by (simp add: one_enat_def) lemmas wary_App\<^sub>h[intro] = wary_App[folded of_nat_lt_hmset_of_enat_iff] lemmas wary_AppE\<^sub>h = wary_AppE[folded of_nat_lt_hmset_of_enat_iff] lemmas wary_num_args_le_arity_head\<^sub>h = wary_num_args_le_arity_head[folded of_nat_le_hmset_of_enat_iff] lemmas wary_apps\<^sub>h = wary_apps[folded of_nat_le_hmset_of_enat_iff] lemmas wary_cases_apps\<^sub>h[consumes 1, case_names apps] = wary_cases_apps[folded of_nat_le_hmset_of_enat_iff] lemmas ground_heads_arity\<^sub>h = ground_heads_arity[folded hmset_of_enat_le] lemmas some_ground_head_arity\<^sub>h = some_ground_head_arity[folded hmset_of_enat_le] lemmas \\<^sub>h_gt_0 = \_gt_0[folded of_nat_less_hmset, unfolded of_nat_0] lemmas \\<^sub>h_le_\\<^sub>h = \_le_\[folded of_nat_le_hmset] lemmas arity_hd\<^sub>h_lt_\_if_\\<^sub>h_gt_0 = arity_hd_ne_infinity_if_\_gt_0 [folded of_nat_less_hmset, unfolded of_nat_0, folded hmset_of_enat_lt_iff_ne_infinity] lemma wt_sym_ge\<^sub>h: "wt_sym f \ \\<^sub>h - \\<^sub>h * arity_sym\<^sub>h f" proof - have "of_nat (the_enat (of_nat \ * arity_sym f)) = \\<^sub>h * arity_sym\<^sub>h f" by (cases "arity_sym f", simp add: of_nat_eq_enat, metis arity_sym_ne_infinity_if_\_gt_0 gr_zeroI mult_eq_0_iff of_nat_0 the_enat_0) thus ?thesis using wt_sym_ge[unfolded of_nat_minus_hmset] by metis qed lemmas unary_wt_sym_0_gt\<^sub>h = unary_wt_sym_0_gt[folded hmset_of_enat_inject, unfolded hmset_of_enat_1] lemmas unary_wt_sym_0_imp_\\<^sub>h_eq_\\<^sub>h = unary_wt_sym_0_imp_\_eq_\ [folded of_nat_inject_hmset, unfolded of_nat_0] lemmas extf_ext_snoc_if_\\<^sub>h_eq_\\<^sub>h = extf_ext_snoc_if_\_eq_\[folded of_nat_inject_hmset] lemmas extf_snoc_if_\\<^sub>h_eq_\\<^sub>h = ext_snoc.snoc[OF extf_ext_snoc_if_\\<^sub>h_eq_\\<^sub>h] lemmas arity_sym\<^sub>h_lt_\_if_\\<^sub>h_gt_0 = arity_sym_ne_infinity_if_\_gt_0 [folded of_nat_less_hmset hmset_of_enat_lt_iff_ne_infinity, unfolded of_nat_0] lemmas arity_var\<^sub>h_lt_\_if_\\<^sub>h_gt_0 = arity_var_ne_infinity_if_\_gt_0 [folded of_nat_less_hmset hmset_of_enat_lt_iff_ne_infinity, unfolded of_nat_0] lemmas arity\<^sub>h_lt_\_if_\\<^sub>h_gt_0 = arity_ne_infinity_if_\_gt_0 [folded of_nat_less_hmset hmset_of_enat_lt_iff_ne_infinity, unfolded of_nat_0] lemmas warywary_subst_subst\<^sub>h_conv = wary_subst_def[folded hmset_of_enat_le] lemmas extf_singleton_nil_if_\\<^sub>h_eq_\\<^sub>h = extf_singleton_nil_if_\_eq_\[folded of_nat_inject_hmset] lemma arity_sym\<^sub>h_if_\\<^sub>h_gt_0_E: assumes \_gt_0: "\\<^sub>h > 0" obtains n where "arity_sym\<^sub>h f = of_nat n" using arity_sym\<^sub>h_lt_\_if_\\<^sub>h_gt_0 assms lt_\_imp_ex_of_nat by blast lemma arity_var\<^sub>h_if_\\<^sub>h_gt_0_E: assumes \_gt_0: "\\<^sub>h > 0" obtains n where "arity_var\<^sub>h f = of_nat n" using arity_var\<^sub>h_lt_\_if_\\<^sub>h_gt_0 assms lt_\_imp_ex_of_nat by blast subsection \Weights and Subterm Coefficients\ abbreviation zhmset_of_tpoly :: "('a, hmultiset) tpoly \ ('a, zhmultiset) tpoly" where "zhmset_of_tpoly \ map_tpoly (\x. x) zhmset_of" abbreviation eval_ztpoly :: "('a \ zhmultiset) \ ('a, hmultiset) tpoly \ zhmultiset" where "eval_ztpoly A p \ eval_tpoly A (zhmset_of_tpoly p)" lemma eval_tpoly_eq_eval_ztpoly[simp]: "zhmset_of (eval_tpoly A p) = eval_ztpoly (\v. zhmset_of (A v)) p" by (induct p, simp_all add: zhmset_of_sum_list zhmset_of_prod_list o_def, simp_all cong: map_cong) definition min_ground_head :: "('s, 'v) hd \ 's" where "min_ground_head \ = (SOME f. f \ ground_heads \ \ (\g \ ground_heads \. wt_sym g + \\<^sub>h * arity_sym\<^sub>h g \ wt_sym f + \\<^sub>h * arity_sym\<^sub>h f))" datatype 'va pvar = PWt 'va | PCoef 'va nat primrec min_passign :: "'v pvar \ hmultiset" where "min_passign (PWt x) = wt_sym (min_ground_head (Var x))" | "min_passign (PCoef _ _) = 1" abbreviation min_zpassign :: "'v pvar \ zhmultiset" where "min_zpassign v \ zhmset_of (min_passign v)" lemma min_zpassign_simps[simp]: "min_zpassign (PWt x) = zhmset_of (wt_sym (min_ground_head (Var x)))" "min_zpassign (PCoef x i) = 1" by (simp_all add: zhmset_of_1) definition legal_passign :: "('v pvar \ hmultiset) \ bool" where "legal_passign A \ (\x. A x \ min_passign x)" definition legal_zpassign :: "('v pvar \ zhmultiset) \ bool" where "legal_zpassign A \ (\x. A x \ min_zpassign x)" lemma legal_min_passign: "legal_passign min_passign" unfolding legal_passign_def by simp lemma legal_min_zpassign: "legal_zpassign min_zpassign" unfolding legal_zpassign_def by simp lemma assign_ge_0[intro]: "legal_zpassign A \ A x \ 0" unfolding legal_zpassign_def by (auto intro: dual_order.trans) definition eq_tpoly :: "('v pvar, hmultiset) tpoly \ ('v pvar, hmultiset) tpoly \ bool" (infix "=\<^sub>p" 50) where "q =\<^sub>p p \ (\A. legal_zpassign A \ eval_ztpoly A q = eval_ztpoly A p)" definition ge_tpoly :: "('v pvar, hmultiset) tpoly \ ('v pvar, hmultiset) tpoly \ bool" (infix "\\<^sub>p" 50) where "q \\<^sub>p p \ (\A. legal_zpassign A \ eval_ztpoly A q \ eval_ztpoly A p)" definition gt_tpoly :: "('v pvar, hmultiset) tpoly \ ('v pvar, hmultiset) tpoly \ bool" (infix ">\<^sub>p" 50) where "q >\<^sub>p p \ (\A. legal_zpassign A \ eval_ztpoly A q > eval_ztpoly A p)" lemma gt_tpoly_imp_ge[intro]: "q >\<^sub>p p \ q \\<^sub>p p" unfolding ge_tpoly_def gt_tpoly_def by (simp add: le_less) lemma eq_tpoly_refl[simp]: "p =\<^sub>p p" unfolding eq_tpoly_def by simp lemma ge_tpoly_refl[simp]: "p \\<^sub>p p" unfolding ge_tpoly_def by simp lemma gt_tpoly_irrefl: "\ p >\<^sub>p p" unfolding gt_tpoly_def legal_zpassign_def by fast lemma eq_eq_tpoly_trans: "r =\<^sub>p q \ q =\<^sub>p p \ r =\<^sub>p p" and eq_ge_tpoly_trans: "r =\<^sub>p q \ q \\<^sub>p p \ r \\<^sub>p p" and eq_gt_tpoly_trans: "r =\<^sub>p q \ q >\<^sub>p p \ r >\<^sub>p p" and ge_eq_tpoly_trans: "r \\<^sub>p q \ q =\<^sub>p p \ r \\<^sub>p p" and ge_ge_tpoly_trans: "r \\<^sub>p q \ q \\<^sub>p p \ r \\<^sub>p p" and ge_gt_tpoly_trans: "r \\<^sub>p q \ q >\<^sub>p p \ r >\<^sub>p p" and gt_eq_tpoly_trans: "r >\<^sub>p q \ q =\<^sub>p p \ r >\<^sub>p p" and gt_ge_tpoly_trans: "r >\<^sub>p q \ q \\<^sub>p p \ r >\<^sub>p p" and gt_gt_tpoly_trans: "r >\<^sub>p q \ q >\<^sub>p p \ r >\<^sub>p p" unfolding eq_tpoly_def ge_tpoly_def gt_tpoly_def by (auto intro: order.trans less_trans less_le_trans le_less_trans)+ primrec coef_hd :: "('s, 'v) hd \ nat \ ('v pvar, hmultiset) tpoly" where "coef_hd (Var x) i = PVar (PCoef x i)" | "coef_hd (Sym f) i = PNum (coef_sym f i)" lemma coef_hd_gt_0: assumes legal: "legal_zpassign A" shows "eval_ztpoly A (coef_hd \ i) > 0" unfolding legal_zpassign_def proof (cases \) case (Var x1) thus ?thesis using legal[unfolded legal_zpassign_def, rule_format, of "PCoef x i" for x] by (auto simp: coef_sym_gt_0 zhmset_of_1 intro: dual_order.strict_trans1 zero_less_one) next case (Sym x2) thus ?thesis using legal[unfolded legal_zpassign_def, rule_format, of "PWt x" for x] by simp (metis coef_sym_gt_0 zhmset_of_0 zhmset_of_less) qed primrec coef :: "('s, 'v) tm \ nat \ ('v pvar, hmultiset) tpoly" where "coef (Hd \) i = coef_hd \ i" | "coef (App s _) i = coef s (i + 1)" lemma coef_apps[simp]: "coef (apps s ss) i = coef s (i + length ss)" by (induct ss arbitrary: s i) auto lemma coef_gt_0: "legal_zpassign A \ eval_ztpoly A (coef s i) > 0" by (induct s arbitrary: i) (auto intro: coef_hd_gt_0) lemma exists_min_ground_head: "\f. f \ ground_heads \ \ (\g \ ground_heads \. wt_sym g + \\<^sub>h * arity_sym\<^sub>h g \ wt_sym f + \\<^sub>h * arity_sym\<^sub>h f)" proof - let ?R = "{(f, g). f \ ground_heads \ \ g \ ground_heads \ \ wt_sym g + \\<^sub>h * arity_sym\<^sub>h g > wt_sym f + \\<^sub>h * arity_sym\<^sub>h f}" have wf_R: "wf ?R" using wf_app[of "{(M, N). M < N}" "\f. wt_sym f + \\<^sub>h * arity_sym\<^sub>h f", OF wf] by (auto intro: wf_subset) have "\f. f \ ground_heads \" by (meson ground_heads_nonempty subsetI subset_empty) thus ?thesis using wf_eq_minimal[THEN iffD1, OF wf_R] by force qed lemma min_ground_head_Sym[simp]: "min_ground_head (Sym f) = f" unfolding min_ground_head_def by auto lemma min_ground_head_in_ground_heads: "min_ground_head \ \ ground_heads \" unfolding min_ground_head_def using someI_ex[OF exists_min_ground_head] by blast lemma min_ground_head_min: "f \ ground_heads \ \ wt_sym f + \\<^sub>h * arity_sym\<^sub>h f \ wt_sym (min_ground_head \) + \\<^sub>h * arity_sym\<^sub>h (min_ground_head \)" unfolding min_ground_head_def using someI_ex[OF exists_min_ground_head] by blast lemma min_ground_head_antimono: "ground_heads \ \ ground_heads \ \ wt_sym (min_ground_head \) + \\<^sub>h * arity_sym\<^sub>h (min_ground_head \) \ wt_sym (min_ground_head \) + \\<^sub>h * arity_sym\<^sub>h (min_ground_head \)" using min_ground_head_in_ground_heads min_ground_head_min by blast primrec wt0 :: "('s, 'v) hd \ ('v pvar, hmultiset) tpoly" where "wt0 (Var x) = PVar (PWt x)" | "wt0 (Sym f) = PNum (wt_sym f)" lemma wt0_ge_min_ground_head: "legal_zpassign A \ eval_ztpoly A (wt0 \) \ zhmset_of (wt_sym (min_ground_head \))" by (cases \, simp_all, metis legal_zpassign_def min_zpassign_simps(1)) lemma eval_ztpoly_nonneg: "legal_zpassign A \ eval_ztpoly A p \ 0" by (induct p) (auto cong: map_cong intro!: sum_list_nonneg prod_list_nonneg) lemma in_zip_imp_size_lt_apps: "(s, y) \ set (zip ss ys) \ size s < size (apps (Hd \) ss)" by (auto dest!: set_zip_leftD simp: size_in_args) function wt :: "('s, 'v) tm \ ('v pvar, hmultiset) tpoly" where "wt (apps (Hd \) ss) = PSum ([wt0 \, PNum (\\<^sub>h * (arity_sym\<^sub>h (min_ground_head \) - of_nat (length ss)))] @ map (\(s, i). PMult [coef_hd \ i, wt s]) (zip ss [0.. ('v pvar \ zhmultiset) \ ('s, 'v) hd \ ('s, 'v) tm list \ zhmultiset" where "wt_args i A \ ss = sum_list (map (eval_ztpoly A \ (\(s, i). PMult [coef_hd \ i, wt s])) (zip ss [i..) = PSum [wt0 \, PNum (\\<^sub>h * arity_sym\<^sub>h (min_ground_head \))]" by (rule wt.simps[of _ "[]", simplified]) lemma coef_hd_cong: "(\x \ vars_hd \. \i. A (PCoef x i) = B (PCoef x i)) \ eval_ztpoly A (coef_hd \ i) = eval_ztpoly B (coef_hd \ i)" by (cases \) auto lemma wt0_cong: assumes pwt_eq: "\x \ vars_hd \. A (PWt x) = B (PWt x)" shows "eval_ztpoly A (wt0 \) = eval_ztpoly B (wt0 \)" using pwt_eq by (cases \) auto lemma wt_cong: assumes "\x \ vars s. A (PWt x) = B (PWt x)" and "\x \ vars s. \i. A (PCoef x i) = B (PCoef x i)" shows "eval_ztpoly A (wt s) = eval_ztpoly B (wt s)" using assms proof (induct s rule: tm_induct_apps) case (apps \ ss) note ih = this(1) and pwt_eq = this(2) and pcoef_eq = this(3) have ih': "eval_ztpoly A (wt s) = eval_ztpoly B (wt s)" if s_in: "s \ set ss" for s proof (rule ih[OF s_in]) show "\x \ vars s. A (PWt x) = B (PWt x)" using pwt_eq s_in by force show "\x \ vars s. \i. A (PCoef x i) = B (PCoef x i)" using pcoef_eq s_in by force qed have wt0_eq: "eval_ztpoly A (wt0 \) = eval_ztpoly B (wt0 \)" by (rule wt0_cong) (simp add: pwt_eq) have coef_\_eq: "eval_ztpoly A (coef_hd \ i) = eval_ztpoly B (coef_hd \ i)" for i by (rule coef_hd_cong) (simp add: pcoef_eq) show ?case using ih' wt0_eq coef_\_eq by (auto dest!: set_zip_leftD intro!: arg_cong[of _ _ sum_list]) qed lemma ground_eval_ztpoly_wt_eq: "ground s \ eval_ztpoly A (wt s) = eval_ztpoly B (wt s)" by (rule wt_cong) auto lemma exists_wt_sym: assumes legal: "legal_zpassign A" shows "\f \ ground_heads \. eval_ztpoly A (wt (Hd \)) \ zhmset_of (wt_sym f + \\<^sub>h * arity_sym\<^sub>h f)" unfolding eq_tpoly_def proof (cases \) case Var thus ?thesis using legal[unfolded legal_zpassign_def] by simp (metis add_le_cancel_right ground_heads.simps(1) min_ground_head_in_ground_heads min_zpassign_simps(1) zhmset_of_plus) next case Sym thus ?thesis by (simp add: zhmset_of_plus) qed lemma wt_ge_\\<^sub>h: assumes legal: "legal_zpassign A" shows "eval_ztpoly A (wt s) \ zhmset_of \\<^sub>h" proof (induct s rule: tm_induct_apps) case (apps \ ss) note ih = this(1) { assume ss_eq_nil: "ss = []" have "\\<^sub>h \ wt_sym (min_ground_head \) + \\<^sub>h * arity_sym\<^sub>h (min_ground_head \)" using wt_sym_ge\<^sub>h[of "min_ground_head \"] by (metis add_diff_cancel_left' leD leI le_imp_minus_plus_hmset le_minus_plus_same_hmset less_le_trans) hence "zhmset_of \\<^sub>h \ zhmset_of (wt_sym (min_ground_head \)) + zhmset_of (\\<^sub>h * arity_sym\<^sub>h (min_ground_head \))" by (metis zhmset_of_le zhmset_of_plus) also have "\ \ eval_tpoly A (map_tpoly (\x. x) zhmset_of (wt0 \)) + zhmset_of (\\<^sub>h * arity_sym\<^sub>h (min_ground_head \))" using wt0_ge_min_ground_head[OF legal] by simp finally have ?case using ss_eq_nil by simp } moreover { let ?arg_wt = "eval_tpoly A \ (map_tpoly (\x. x) zhmset_of \ (\(s, i). PMult [coef_hd \ i, wt s]))" assume ss_ne_nil: "ss \ []" hence "zhmset_of \\<^sub>h \ eval_tpoly A (map_tpoly (\x. x) zhmset_of (PMult [coef_hd \ 0, wt (hd ss)]))" by (simp add: ih coef_hd_gt_0[OF legal] nonneg_le_mult_right_mono_zhmset) also have "\ = hd (map ?arg_wt (zip ss [0.. \ sum_list (map ?arg_wt (zip ss [0.. \ eval_tpoly A (map_tpoly (\x. x) zhmset_of (wt0 \)) + (zhmset_of (\\<^sub>h * (arity_sym\<^sub>h (min_ground_head \) - of_nat (length ss))) + sum_list (map ?arg_wt (zip ss [0.. eval_tpoly A (map_tpoly (\p. p) zhmset_of (wt0 \))" using legal eval_ztpoly_nonneg by blast then show ?thesis by (meson leD leI le_add_same_cancel2 less_le_trans zhmset_of_nonneg) qed finally have ?case by simp } ultimately show ?case by linarith qed lemma wt_args_ge_length_times_\\<^sub>h: assumes legal: "legal_zpassign A" shows "wt_args i A \ ss \ of_nat (length ss) * zhmset_of \\<^sub>h" unfolding wt_args_def by (rule sum_list_ge_length_times[unfolded wt_args_def, of "map (eval_ztpoly A \ (\(s, i). PMult [coef_hd \ i, wt s])) (zip ss [i..\<^sub>h) lemma wt_ge_\\<^sub>h: "legal_zpassign A \ eval_ztpoly A (wt s) \ zhmset_of \\<^sub>h" using \\<^sub>h_le_\\<^sub>h[folded zhmset_of_le] order.trans wt_ge_\\<^sub>h zhmset_of_le by blast lemma wt_gt_0: "legal_zpassign A \ eval_ztpoly A (wt s) > 0" using \\<^sub>h_gt_0[folded zhmset_of_less, unfolded zhmset_of_0] wt_ge_\\<^sub>h by (blast intro: less_le_trans) lemma wt_gt_\\<^sub>h_if_superunary: assumes legal: "legal_zpassign A" and superunary: "arity_hd\<^sub>h (head s) > 1" shows "eval_ztpoly A (wt s) > zhmset_of \\<^sub>h" proof (cases "\\<^sub>h = \\<^sub>h") case \_ne_\: False show ?thesis using order.not_eq_order_implies_strict[OF \_ne_\ \\<^sub>h_le_\\<^sub>h, folded zhmset_of_less] wt_ge_\\<^sub>h[OF legal] by (blast intro: less_le_trans) next case \_eq_\: True show ?thesis using superunary proof (induct s rule: tm_induct_apps) case (apps \ ss) have "arity_hd\<^sub>h \ > 1" using apps(2) by simp hence min_gr_ary: "arity_sym\<^sub>h (min_ground_head \) > 1" using ground_heads_arity\<^sub>h less_le_trans min_ground_head_in_ground_heads by blast have "zhmset_of \\<^sub>h < eval_ztpoly A (wt0 \) + zhmset_of (\\<^sub>h * arity_sym\<^sub>h (min_ground_head \))" unfolding \_eq_\ by (rule add_strict_increasing2[OF eval_ztpoly_nonneg[OF legal]], unfold zhmset_of_less, rule gt_0_lt_mult_gt_1_hmset[OF \\<^sub>h_gt_0 min_gr_ary]) also have "\ \ eval_ztpoly A (wt0 \) + zhmset_of (\\<^sub>h * (arity_sym\<^sub>h (min_ground_head \) - of_nat (length ss))) + zhmset_of (of_nat (length ss) * \\<^sub>h)" by (auto simp: \\<^sub>h_gt_0 \_eq_\ zmset_of_le zhmset_of_plus[symmetric] algebra_simps simp del: ring_distribs simp: ring_distribs[symmetric]) (metis add.commute le_minus_plus_same_hmset) also have "\ \ eval_ztpoly A (wt0 \) + zhmset_of (\\<^sub>h * (arity_sym\<^sub>h (min_ground_head \) - of_nat (length ss))) + wt_args 0 A \ ss" using wt_args_ge_length_times_\\<^sub>h[OF legal] by (simp add: zhmset_of_times of_nat_zhmset) finally show ?case by (simp add: wt_args_def add_ac(1) comp_def) qed qed lemma wt_App_plus_\\<^sub>h_ge: "eval_ztpoly A (wt (App s t)) + zhmset_of \\<^sub>h \ eval_ztpoly A (wt s) + eval_ztpoly A (coef s 0) * eval_ztpoly A (wt t)" proof (cases s rule: tm_exhaust_apps) case s: (apps \ ss) show ?thesis proof (cases "arity_sym\<^sub>h (min_ground_head \) = \") case ary_eq_\: True show ?thesis unfolding ary_eq_\ s App_apps wt.simps by (auto simp: diff_diff_add_hmset[symmetric] add.assoc) next case False show ?thesis unfolding s App_apps wt.simps by (simp add: algebra_simps zhmset_of_plus[symmetric] zmset_of_le, simp del: diff_diff_add_hmset add: add.commute[of 1] le_minus_plus_same_hmset distrib_left[of _ "1 :: hmultiset", unfolded mult.right_neutral, symmetric] diff_diff_add_hmset[symmetric]) qed qed lemma wt_App_fun_\\<^sub>h: assumes legal: "legal_zpassign A" and wt_st: "eval_ztpoly A (wt (App s t)) = eval_ztpoly A (wt t)" shows "eval_ztpoly A (wt s) = zhmset_of \\<^sub>h" proof - have "eval_ztpoly A (wt (App s t)) = eval_ztpoly A (wt t)" using wt_st by simp hence wt_s_t_le_\_t: "eval_ztpoly A (wt s) + eval_ztpoly A (coef s 0) * eval_ztpoly A (wt t) \ zhmset_of \\<^sub>h + eval_ztpoly A (wt t)" using wt_App_plus_\\<^sub>h_ge by (metis add.commute) also have "\ \ eval_ztpoly A (wt s) + eval_ztpoly A (wt t)" using wt_ge_\\<^sub>h[OF legal] by simp finally have "eval_ztpoly A (coef s 0) * eval_ztpoly A (wt t) \ eval_ztpoly A (wt t)" by simp hence "eval_ztpoly A (coef s 0) = 1" using eval_ztpoly_nonneg[OF legal] by (metis (no_types, lifting) coef_gt_0 dual_order.order_iff_strict leD legal mult_cancel_right1 nonneg_le_mult_right_mono_zhmset wt_gt_0) thus ?thesis using wt_s_t_le_\_t by (simp add: add.commute antisym wt_ge_\\<^sub>h[OF legal]) qed lemma wt_App_arg_\\<^sub>h: assumes legal: "legal_zpassign A" and wt_st: "eval_ztpoly A (wt (App s t)) = eval_ztpoly A (wt s)" shows "eval_ztpoly A (wt t) = zhmset_of \\<^sub>h" proof - have "eval_ztpoly A (wt (App s t)) + zhmset_of \\<^sub>h = eval_ztpoly A (wt s) + zhmset_of \\<^sub>h" using wt_st by simp hence "eval_ztpoly A (coef s 0) * eval_ztpoly A (wt t) \ zhmset_of \\<^sub>h" (is "?k * ?w \ _") by (metis add_le_cancel_left wt_App_plus_\\<^sub>h_ge) hence "?k * ?w = zhmset_of \\<^sub>h" using wt_ge_\\<^sub>h[OF legal] coef_gt_0[OF legal, unfolded zero_less_iff_1_le_hmset] by (simp add: antisym nonneg_le_mult_right_mono_zhmset) hence "?w \ zhmset_of \\<^sub>h" by (metis coef_gt_0[OF legal] dual_order.order_iff_strict eval_ztpoly_nonneg[OF legal] nonneg_le_mult_right_mono_zhmset) thus ?thesis by (simp add: antisym wt_ge_\\<^sub>h[OF legal]) qed lemma wt_App_ge_fun: "wt (App s t) \\<^sub>p wt s" unfolding ge_tpoly_def proof clarify fix A assume legal: "legal_zpassign A" have "zhmset_of \\<^sub>h \ eval_ztpoly A (coef s 0) * eval_ztpoly A (wt t)" by (simp add: coef_gt_0 legal nonneg_le_mult_right_mono_zhmset wt_ge_\\<^sub>h) hence "eval_ztpoly A (wt s) + zhmset_of \\<^sub>h \ eval_ztpoly A (wt (App s t)) + zhmset_of \\<^sub>h" by (metis add_le_cancel_right add_less_le_mono not_le wt_App_plus_\\<^sub>h_ge) thus "eval_ztpoly A (wt s) \ eval_ztpoly A (wt (App s t))" by simp qed lemma wt_App_ge_arg: "wt (App s t) \\<^sub>p wt t" unfolding ge_tpoly_def by (cases s rule: tm_exhaust_apps, simp, unfold App_apps wt.simps) (auto simp: comp_def coef_hd_gt_0 eval_ztpoly_nonneg nonneg_le_mult_right_mono_zhmset intro!: sum_list_nonneg eval_ztpoly_nonneg add_increasing) lemma wt_\\<^sub>h_imp_\\<^sub>h_eq_\\<^sub>h: assumes legal: "legal_zpassign A" and wt_s_eq_\: "eval_ztpoly A (wt s) = zhmset_of \\<^sub>h" shows "\\<^sub>h = \\<^sub>h" using \\<^sub>h_le_\\<^sub>h wt_ge_\\<^sub>h[OF legal, of s, unfolded wt_s_eq_\ zhmset_of_le] by (rule antisym) lemma wt_ge_vars: "wt t \\<^sub>p wt s \ vars t \ vars s" proof (induct s) case t: (Hd \) note wt_ge_\ = this(1) show ?case proof (cases \) case \: (Var x) { assume z_ni_t: "x \ vars t" let ?A = min_zpassign let ?B = "\v. if v = PWt x then eval_ztpoly ?A (wt t) + ?A v + 1 else ?A v" have legal_B: "legal_zpassign ?B" unfolding legal_zpassign_def by (auto simp: legal_min_zpassign intro!: add_increasing eval_ztpoly_nonneg) have eval_B_eq_A: "eval_ztpoly ?B (wt t) = eval_ztpoly ?A (wt t)" by (rule wt_cong) (auto simp: z_ni_t) have "eval_ztpoly ?B (wt (Hd (Var x))) > eval_ztpoly ?B (wt t)" by (auto simp: eval_B_eq_A zero_less_iff_1_le_zhmset zhmset_of_plus[symmetric] algebra_simps) hence False using wt_ge_\ \ unfolding ge_tpoly_def by (blast dest: leD intro: legal_B legal_min_zpassign) } thus ?thesis by (auto simp: \) qed simp next case (App s1 s2) note ih1 = this(1) and ih2 = this(2) and wt_t_ge_wt_s1s2 = this(3) have "vars s1 \ vars t" using ih1 wt_t_ge_wt_s1s2 wt_App_ge_fun order_trans unfolding ge_tpoly_def by blast moreover have "vars s2 \ vars t" using ih2 wt_t_ge_wt_s1s2 wt_App_ge_arg order_trans unfolding ge_tpoly_def by blast ultimately show ?case by simp qed lemma sum_coefs_ge_num_args_if_\\<^sub>h_eq_0: assumes legal: "legal_passign A" and \_eq_0: "\\<^sub>h = 0" and wary_s: "wary s" shows "sum_coefs (eval_tpoly A (wt s)) \ num_args s" proof (cases s rule: tm_exhaust_apps) case s: (apps \ ss) show ?thesis unfolding s proof (induct ss rule: rev_induct) case (snoc sa ss) note ih = this let ?Az = "\v. zhmset_of (A v)" have legalz: "legal_zpassign ?Az" using legal unfolding legal_passign_def legal_zpassign_def zhmset_of_le by assumption have "eval_ztpoly ?Az (coef_hd \ (length ss)) > 0" using legal coef_hd_gt_0 eval_tpoly_eq_eval_ztpoly by (simp add: coef_hd_gt_0[OF legalz]) hence k: "eval_tpoly A (coef_hd \ (length ss)) > 0" (is "?k > _") unfolding eval_tpoly_eq_eval_ztpoly[symmetric] zhmset_of_less[symmetric] zhmset_of_0 by assumption have "eval_ztpoly ?Az (wt sa) > 0" (is "?w > _") by (simp add: wt_gt_0[OF legalz]) hence w: "eval_tpoly A (wt sa) > 0" (is "?w > _") unfolding eval_tpoly_eq_eval_ztpoly[symmetric] zhmset_of_less[symmetric] zhmset_of_0 by assumption have "?k * ?w > 0" using k w by simp hence "sum_coefs (?k * ?w) > 0" by (rule sum_coefs_gt_0[THEN iffD2]) thus ?case using ih by (simp del: apps_append add: s \_eq_0) qed simp qed subsection \Inductive Definitions\ inductive gt :: "('s, 'v) tm \ ('s, 'v) tm \ bool" (infix ">\<^sub>t" 50) where gt_wt: "wt t >\<^sub>p wt s \ t >\<^sub>t s" | gt_unary: "wt t \\<^sub>p wt s \ \ head t \\\<^sub>h\<^sub>d head s \ num_args t = 1 \ (\f \ ground_heads (head t). arity_sym f = 1 \ wt_sym f = 0) \ arg t >\<^sub>t s \ arg t = s \ t >\<^sub>t s" | gt_diff: "wt t \\<^sub>p wt s \ head t >\<^sub>h\<^sub>d head s \ t >\<^sub>t s" | gt_same: "wt t \\<^sub>p wt s \ head t = head s \ (\f \ ground_heads (head t). extf f (>\<^sub>t) (args t) (args s)) \ t >\<^sub>t s" abbreviation ge :: "('s, 'v) tm \ ('s, 'v) tm \ bool" (infix "\\<^sub>t" 50) where "t \\<^sub>t s \ t >\<^sub>t s \ t = s" inductive gt_wt :: "('s, 'v) tm \ ('s, 'v) tm \ bool" where gt_wtI: "wt t >\<^sub>p wt s \ gt_wt t s" inductive gt_unary :: "('s, 'v) tm \ ('s, 'v) tm \ bool" where gt_unaryI: "wt t \\<^sub>p wt s \ \ head t \\\<^sub>h\<^sub>d head s \ num_args t = 1 \ (\f \ ground_heads (head t). arity_sym f = 1 \ wt_sym f = 0) \ arg t \\<^sub>t s \ gt_unary t s" inductive gt_diff :: "('s, 'v) tm \ ('s, 'v) tm \ bool" where gt_diffI: "wt t \\<^sub>p wt s \ head t >\<^sub>h\<^sub>d head s \ gt_diff t s" inductive gt_same :: "('s, 'v) tm \ ('s, 'v) tm \ bool" where gt_sameI: "wt t \\<^sub>p wt s \ head t = head s \ (\f \ ground_heads (head t). extf f (>\<^sub>t) (args t) (args s)) \ gt_same t s" lemma gt_iff_wt_unary_diff_same: "t >\<^sub>t s \ gt_wt t s \ gt_unary t s \ gt_diff t s \ gt_same t s" by (subst gt.simps) (auto simp: gt_wt.simps gt_unary.simps gt_diff.simps gt_same.simps) lemma gt_imp_wt: "t >\<^sub>t s \ wt t \\<^sub>p wt s" by (blast elim: gt.cases) lemma gt_imp_vars: "t >\<^sub>t s \ vars t \ vars s" by (erule wt_ge_vars[OF gt_imp_wt]) subsection \Irreflexivity\ theorem gt_irrefl: "wary s \ \ s >\<^sub>t s" proof (induct "size s" arbitrary: s rule: less_induct) case less note ih = this(1) and wary_s = this(2) show ?case proof assume s_gt_s: "s >\<^sub>t s" show False using s_gt_s proof (cases rule: gt.cases) case gt_same then obtain f where f: "extf f (>\<^sub>t) (args s) (args s)" by fastforce thus False using wary_s ih by (metis wary_args extf_irrefl size_in_args) qed (auto simp: comp_hd_def gt_tpoly_irrefl gt_hd_irrefl) qed qed subsection \Transitivity\ lemma not_extf_gt_nil_singleton_if_\\<^sub>h_eq_\\<^sub>h: assumes wary_s: "wary s" and \_eq_\: "\\<^sub>h = \\<^sub>h" shows "\ extf f (>\<^sub>t) [] [s]" proof assume nil_gt_s: "extf f (>\<^sub>t) [] [s]" note s_gt_nil = extf_singleton_nil_if_\\<^sub>h_eq_\\<^sub>h[OF \_eq_\, of f gt s] have "\ extf f (>\<^sub>t) [] []" by (rule extf_irrefl) simp moreover have "extf f (>\<^sub>t) [] []" using extf_trans_from_irrefl[of "{s}", OF _ _ _ _ _ _ nil_gt_s s_gt_nil] gt_irrefl[OF wary_s] by fastforce ultimately show False by sat qed lemma gt_sub_arg: "wary (App s t) \ App s t >\<^sub>t t" proof (induct t arbitrary: s rule: measure_induct_rule[of size]) case (less t) note ih = this(1) and wary_st = this(2) { fix A assume legal: "legal_zpassign A" and wt_st: "eval_ztpoly A (wt (App s t)) = eval_ztpoly A (wt t)" have \_eq_\: "\\<^sub>h = \\<^sub>h" using wt_App_fun_\\<^sub>h[OF legal] wt_\\<^sub>h_imp_\\<^sub>h_eq_\\<^sub>h[OF legal] wt_st by blast hence \_gt_0: "\\<^sub>h > 0" using \\<^sub>h_gt_0 by simp have wt_s: "eval_ztpoly A (wt s) = zhmset_of \\<^sub>h" by (rule wt_App_fun_\\<^sub>h[OF legal wt_st]) have wary_t: "wary t" by (rule wary_AppE\<^sub>h[OF wary_st]) have nargs_lt: "of_nat (num_args s) < arity_hd\<^sub>h (head s)" by (rule wary_AppE\<^sub>h[OF wary_st]) have ary_hd_s: "arity_hd\<^sub>h (head s) = 1" by (metis gr_implies_not_zero_hmset legal lt_1_iff_eq_0_hmset nargs_lt neq_iff wt_gt_\\<^sub>h_if_superunary wt_s) hence nargs_s: "num_args s = 0" by (metis less_one nargs_lt of_nat_1 of_nat_less_hmset) hence s_eq_hd: "s = Hd (head s)" by (simp add: Hd_head_id) obtain f where f_in: "f \ ground_heads (head s)" and wt_f_etc: "wt_sym f + \\<^sub>h * arity_sym\<^sub>h f = \\<^sub>h" proof - assume a: "\f. \f \ local.ground_heads (head s); wt_sym f + \\<^sub>h * arity_sym\<^sub>h f = \\<^sub>h\ \ thesis" have "\f. \\<^sub>h - \\<^sub>h * arity_sym\<^sub>h f \ wt_sym f" using wt_s by (metis legal wt_\\<^sub>h_imp_\\<^sub>h_eq_\\<^sub>h wt_sym_ge\<^sub>h) hence "\s. \ \\<^sub>h * arity_sym\<^sub>h s + wt_sym s < \\<^sub>h" by (metis add_diff_cancel_left' le_imp_minus_plus_hmset leD le_minus_plus_same_hmset less_le_trans) thus thesis using a wt_s s_eq_hd by (metis exists_wt_sym legal add.commute order.not_eq_order_implies_strict zhmset_of_le) qed have ary_f_1: "arity_sym f = 1" by (metis \_gt_0 add_diff_cancel_left' ary_hd_s diff_le_self_hmset dual_order.order_iff_strict f_in ground_heads_arity\<^sub>h gt_0_lt_mult_gt_1_hmset hmset_of_enat_1 hmset_of_enat_inject leD wt_f_etc) hence wt_f_0: "wt_sym f = 0" using wt_f_etc by simp { assume hd_s_ncmp_t: "\ head s \\\<^sub>h\<^sub>d head t" have ?case by (rule gt_unary[OF wt_App_ge_arg]) (auto simp: hd_s_ncmp_t nargs_s intro: f_in ary_f_1 wt_f_0) } moreover { assume hd_s_gt_t: "head s >\<^sub>h\<^sub>d head t" have ?case by (rule gt_diff[OF wt_App_ge_arg]) (simp add: hd_s_gt_t) } moreover { assume "head t >\<^sub>h\<^sub>d head s" hence False using ary_f_1 wt_f_0 f_in gt_hd_irrefl gt_sym_antisym unary_wt_sym_0_gt\<^sub>h hmset_of_enat_1 unfolding gt_hd_def by metis } moreover { assume hd_t_eq_s: "head t = head s" hence nargs_t_le: "num_args t \ 1" using ary_hd_s wary_num_args_le_arity_head\<^sub>h[OF wary_t] of_nat_le_hmset by fastforce have extf: "extf f (>\<^sub>t) [t] (args t)" for f proof (cases "args t") case Nil thus ?thesis by (simp add: extf_singleton_nil_if_\\<^sub>h_eq_\\<^sub>h[OF \_eq_\]) next case args_t: (Cons ta ts) hence ts: "ts = []" using ary_hd_s[folded hd_t_eq_s] wary_num_args_le_arity_head\<^sub>h[OF wary_t] of_nat_le_hmset nargs_t_le by simp have ta: "ta = arg t" by (metis apps.simps(1) apps.simps(2) args_t tm.sel(6) tm_collapse_apps ts) hence t: "t = App (fun t) ta" by (metis args.simps(1) args_t not_Cons_self2 tm.exhaust_sel ts) have "t >\<^sub>t ta" by (rule ih[of ta "fun t", folded t, OF _ wary_t]) (metis ta size_arg_lt t tm.disc(2)) thus ?thesis unfolding args_t ts by (metis extf_singleton gt_irrefl wary_t) qed have ?case by (rule gt_same[OF wt_App_ge_arg]) (simp_all add: hd_t_eq_s length_0_conv[THEN iffD1, OF nargs_s] extf) } ultimately have ?case unfolding comp_hd_def by metis } thus ?case using gt_wt by (metis ge_tpoly_def gt_tpoly_def wt_App_ge_arg order.not_eq_order_implies_strict) qed lemma gt_arg: "wary s \ is_App s \ s >\<^sub>t arg s" by (cases s) (auto intro: gt_sub_arg) theorem gt_trans: "wary u \ wary t \ wary s \ u >\<^sub>t t \ t >\<^sub>t s \ u >\<^sub>t s" proof (simp only: atomize_imp, rule measure_induct_rule[of "\(u, t, s). {#size u, size t, size s#}" "\(u, t, s). wary u \ wary t \ wary s \ u >\<^sub>t t \ t >\<^sub>t s \ u >\<^sub>t s" "(u, t, s)", simplified prod.case], simp only: split_paired_all prod.case atomize_imp[symmetric]) fix u t s assume ih: "\ua ta sa. {#size ua, size ta, size sa#} < {#size u, size t, size s#} \ wary ua \ wary ta \ wary sa \ ua >\<^sub>t ta \ ta >\<^sub>t sa \ ua >\<^sub>t sa" and wary_u: "wary u" and wary_t: "wary t" and wary_s: "wary s" and u_gt_t: "u >\<^sub>t t" and t_gt_s: "t >\<^sub>t s" have wt_u_ge_t: "wt u \\<^sub>p wt t" and wt_t_ge_s: "wt t \\<^sub>p wt s" using gt_imp_wt u_gt_t t_gt_s by auto hence wt_u_ge_s: "wt u \\<^sub>p wt s" by (rule ge_ge_tpoly_trans) have wary_arg_u: "wary (arg u)" by (rule wary_arg[OF wary_u]) have wary_arg_t: "wary (arg t)" by (rule wary_arg[OF wary_t]) have wary_arg_s: "wary (arg s)" by (rule wary_arg[OF wary_s]) show "u >\<^sub>t s" using t_gt_s proof cases case gt_wt_t_s: gt_wt hence "wt u >\<^sub>p wt s" using wt_u_ge_t ge_gt_tpoly_trans by blast thus ?thesis by (rule gt_wt) next case gt_unary_t_s: gt_unary have t_app: "is_App t" by (metis args_Nil_iff_is_Hd gt_unary_t_s(3) length_greater_0_conv less_numeral_extra(1)) hence nargs_fun_t: "num_args (fun t) < arity_hd (head (fun t))" by (metis tm.collapse(2) wary_AppE wary_t) have \_eq_\: "\\<^sub>h = \\<^sub>h" using gt_unary_t_s(4) unary_wt_sym_0_imp_\\<^sub>h_eq_\\<^sub>h by blast show ?thesis using u_gt_t proof cases case gt_wt_u_t: gt_wt hence "wt u >\<^sub>p wt s" using wt_t_ge_s gt_ge_tpoly_trans by blast thus ?thesis by (rule gt_wt) next case gt_unary_u_t: gt_unary have u_app: "is_App u" by (metis args_Nil_iff_is_Hd gt_unary_u_t(3) length_greater_0_conv less_numeral_extra(1)) hence nargs_fun_u: "num_args (fun u) = 0" by (metis args.simps(1) gt_unary_u_t(3) list.size(3) one_arg_imp_Hd tm.collapse(2)) have arg_u_gt_s: "arg u >\<^sub>t s" using ih[of "arg u" t s] u_app gt_unary_u_t(5) t_gt_s size_arg_lt wary_arg_u wary_s wary_t by force hence arg_u_ge_s: "arg u \\<^sub>t s" by sat { assume "size (arg u) < size t" hence "{#size u, size (arg u), size s#} < {#size u, size t, size s#}" by simp hence ?thesis using ih[of u "arg u" s] arg_u_gt_s gt_arg u_app wary_s wary_u by blast } moreover { assume "size (arg t) < size s" hence "u >\<^sub>t arg t" using ih[of u t "arg t"] args_Nil_iff_is_Hd gt_arg gt_unary_t_s(3) u_gt_t wary_t wary_u by force hence ?thesis using ih[of u "arg t" s] args_Nil_iff_is_Hd gt_unary_t_s(3,5) size_arg_lt wary_arg_t wary_s wary_u by force } moreover { assume sz_u_gt_t: "size u > size t" and sz_t_gt_s: "size t > size s" { assume hd_u_eq_s: "head u = head s" hence ary_hd_s: "arity_hd (head s) = 1" using ground_heads_arity gt_unary_u_t(3,4) hd_u_eq_s one_enat_def wary_num_args_le_arity_head wary_u by fastforce have extf: "extf f (>\<^sub>t) (args u) (args s)" for f proof (cases "args s") case Nil thus ?thesis by (metis \_eq_\ args.elims args_Nil_iff_is_Hd extf_snoc_if_\\<^sub>h_eq_\\<^sub>h length_0_conv nargs_fun_u tm.sel(4) u_app) next case args_s: (Cons sa ss) hence ss: "ss = []" by (cases s, simp, metis One_nat_def antisym_conv ary_hd_s diff_Suc_1 enat_ord_simps(1) le_add2 length_0_conv length_Cons list.size(4) one_enat_def wary_num_args_le_arity_head wary_s) have sa: "sa = arg s" by (metis apps.simps(1) apps.simps(2) args_s tm.sel(6) tm_collapse_apps ss) have s_app: "is_App s" using args_Nil_iff_is_Hd args_s by force have args_u: "args u = [arg u]" by (metis append_Nil args.simps(2) args_Nil_iff_is_Hd gt_unary_u_t(3) length_0_conv nargs_fun_u tm.collapse(2) zero_neq_one) have max_sz_arg_u_t_arg_t: "Max {size (arg t), size t, size (arg u)} < size u" using size_arg_lt sz_u_gt_t t_app u_app by fastforce have "{#size (arg u), size t, size (arg t)#} < {#size u, size t, size s#}" using max_sz_arg_u_t_arg_t by (auto intro!: Max_lt_imp_lt_mset) hence arg_u_gt_arg_t: "arg u >\<^sub>t arg t" using ih[OF _ wary_arg_u wary_t wary_arg_t] args_Nil_iff_is_Hd gt_arg gt_unary_t_s(3) gt_unary_u_t(5) wary_t by force have max_sz_arg_s_s_arg_t: "Max {size (arg s), size s, size (arg t)} < size u" using s_app t_app size_arg_lt sz_t_gt_s sz_u_gt_t by force have "{#size (arg t), size s, size (arg s)#} < {#size u, size t, size s#}" using max_sz_arg_s_s_arg_t by (auto intro!: Max_lt_imp_lt_mset) hence arg_t_gt_arg_s: "arg t >\<^sub>t arg s" using ih[OF _ wary_arg_t wary_s wary_arg_s] gt_unary_t_s(5) gt_arg args_Nil_iff_is_Hd args_s wary_s by force have "{#size (arg u), size (arg t), size (arg s)#} < {#size u, size t, size s#}" by (auto intro!: add_mset_lt_lt_lt simp: size_arg_lt u_app t_app s_app) hence "arg u >\<^sub>t arg s" using ih[of "arg u" "arg t" "arg s"] arg_u_gt_arg_t arg_t_gt_arg_s wary_arg_s wary_arg_t wary_arg_u by blast thus ?thesis unfolding args_u args_s ss sa by (metis extf_singleton gt_irrefl wary_arg_u) qed have ?thesis by (rule gt_same[OF wt_u_ge_s hd_u_eq_s]) (simp add: extf) } moreover { assume "head u >\<^sub>h\<^sub>d head s" hence ?thesis by (rule gt_diff[OF wt_u_ge_s]) } moreover { assume "head s >\<^sub>h\<^sub>d head u" hence False using gt_hd_def gt_hd_irrefl gt_sym_antisym gt_unary_u_t(4) unary_wt_sym_0_gt by blast } moreover { assume "\ head u \\\<^sub>h\<^sub>d head s" hence ?thesis by (rule gt_unary[OF wt_u_ge_s _ gt_unary_u_t(3,4) arg_u_ge_s]) } ultimately have ?thesis unfolding comp_hd_def by sat } ultimately show ?thesis by (meson less_le_trans linorder_not_le size_arg_lt t_app u_app) next case gt_diff_u_t: gt_diff have False using gt_diff_u_t(2) gt_hd_def gt_hd_irrefl gt_sym_antisym gt_unary_t_s(4) unary_wt_sym_0_gt by blast thus ?thesis by sat next case gt_same_u_t: gt_same have hd_u_ncomp_s: "\ head u \\\<^sub>h\<^sub>d head s" by (rule gt_unary_t_s(2)[folded gt_same_u_t(2)]) have "\f \ ground_heads (head u). arity_sym f = 1 \ wt_sym f = 0" by (rule gt_unary_t_s(4)[folded gt_same_u_t(2)]) hence "arity_hd (head u) = 1" by (metis dual_order.order_iff_strict gr_implies_not_zero_hmset ground_heads_arity gt_same_u_t(2) head_fun hmset_of_enat_1 hmset_of_enat_less lt_1_iff_eq_0_hmset nargs_fun_t) hence "num_args u \ 1" using of_nat_le_hmset wary_num_args_le_arity_head\<^sub>h wary_u by fastforce hence nargs_u: "num_args u = 1" by (cases "args u", metis Hd_head_id \_eq_\ append_Nil args.simps(2) ex_in_conv[THEN iffD2, OF ground_heads_nonempty] gt_same_u_t(2,3) gt_unary_t_s(3) head_fun list.size(3) not_extf_gt_nil_singleton_if_\\<^sub>h_eq_\\<^sub>h one_arg_imp_Hd tm.collapse(2)[OF t_app] wary_arg_t, simp) hence u_app: "is_App u" by (cases u) auto have "arg u >\<^sub>t arg t" by (metis extf_singleton[THEN iffD1] append_Nil args.simps args_Nil_iff_is_Hd comp_hd_def gt_hd_def gt_irrefl gt_same_u_t(2,3) gt_unary_t_s(2,3) head_fun length_0_conv nargs_u one_arg_imp_Hd t_app tm.collapse(2) u_gt_t wary_u) moreover have "{#size (arg u), size (arg t), size s#} < {#size u, size t, size s#}" by (auto intro!: add_mset_lt_lt_lt simp: size_arg_lt u_app t_app) ultimately have "arg u >\<^sub>t s" using ih[OF _ wary_arg_u wary_arg_t wary_s] gt_unary_t_s(5) by blast hence arg_u_ge_s: "arg u \\<^sub>t s" by sat show ?thesis by (rule gt_unary[OF wt_u_ge_s hd_u_ncomp_s nargs_u _ arg_u_ge_s]) (simp add: gt_same_u_t(2) gt_unary_t_s(4)) qed next case gt_diff_t_s: gt_diff show ?thesis using u_gt_t proof cases case gt_wt_u_t: gt_wt hence "wt u >\<^sub>p wt s" using wt_t_ge_s gt_ge_tpoly_trans by blast thus ?thesis by (rule gt_wt) next case gt_unary_u_t: gt_unary have u_app: "is_App u" by (metis args_Nil_iff_is_Hd gt_unary_u_t(3) length_greater_0_conv less_numeral_extra(1)) hence "arg u >\<^sub>t s" using ih[of "arg u" t s] gt_unary_u_t(5) t_gt_s size_arg_lt wary_arg_u wary_s wary_t by force hence arg_u_ge_s: "arg u \\<^sub>t s" by sat { assume "head u = head s" hence False using gt_diff_t_s(2) gt_unary_u_t(2) unfolding comp_hd_def by force } moreover { assume "head s >\<^sub>h\<^sub>d head u" hence False using gt_hd_def gt_hd_irrefl gt_sym_antisym gt_unary_u_t(4) unary_wt_sym_0_gt by blast } moreover { assume "head u >\<^sub>h\<^sub>d head s" hence ?thesis by (rule gt_diff[OF wt_u_ge_s]) } moreover { assume "\ head u \\\<^sub>h\<^sub>d head s" hence ?thesis by (rule gt_unary[OF wt_u_ge_s _ gt_unary_u_t(3,4) arg_u_ge_s]) } ultimately show ?thesis unfolding comp_hd_def by sat next case gt_diff_u_t: gt_diff have "head u >\<^sub>h\<^sub>d head s" using gt_diff_u_t(2) gt_diff_t_s(2) gt_hd_trans by blast thus ?thesis by (rule gt_diff[OF wt_u_ge_s]) next case gt_same_u_t: gt_same have "head u >\<^sub>h\<^sub>d head s" using gt_diff_t_s(2) gt_same_u_t(2) by simp thus ?thesis by (rule gt_diff[OF wt_u_ge_s]) qed next case gt_same_t_s: gt_same show ?thesis using u_gt_t proof cases case gt_wt_u_t: gt_wt hence "wt u >\<^sub>p wt s" using wt_t_ge_s gt_ge_tpoly_trans by blast thus ?thesis by (rule gt_wt) next case gt_unary_u_t: gt_unary have "is_App u" by (metis args_Nil_iff_is_Hd gt_unary_u_t(3) length_greater_0_conv less_numeral_extra(1)) hence "arg u >\<^sub>t s" using ih[of "arg u" t s] gt_unary_u_t(5) t_gt_s size_arg_lt wary_arg_u wary_s wary_t by force hence arg_u_ge_s: "arg u \\<^sub>t s" by sat have "\ head u \\\<^sub>h\<^sub>d head s" using gt_same_t_s(2) gt_unary_u_t(2) by simp thus ?thesis by (rule gt_unary[OF wt_u_ge_s _ gt_unary_u_t(3,4) arg_u_ge_s]) next case gt_diff_u_t: gt_diff have "head u >\<^sub>h\<^sub>d head s" using gt_diff_u_t(2) gt_same_t_s(2) by simp thus ?thesis by (rule gt_diff[OF wt_u_ge_s]) next case gt_same_u_t: gt_same have hd_u_s: "head u = head s" by (simp only: gt_same_t_s(2) gt_same_u_t(2)) let ?S = "set (args u) \ set (args t) \ set (args s)" have gt_trans_args: "\ua \ ?S. \ta \ ?S. \sa \ ?S. ua >\<^sub>t ta \ ta >\<^sub>t sa \ ua >\<^sub>t sa" proof clarify fix sa ta ua assume ua_in: "ua \ ?S" and ta_in: "ta \ ?S" and sa_in: "sa \ ?S" and ua_gt_ta: "ua >\<^sub>t ta" and ta_gt_sa: "ta >\<^sub>t sa" have wary_sa: "wary sa" and wary_ta: "wary ta" and wary_ua: "wary ua" using wary_args ua_in ta_in sa_in wary_u wary_t wary_s by blast+ show "ua >\<^sub>t sa" by (auto intro!: ih[OF Max_lt_imp_lt_mset wary_ua wary_ta wary_sa ua_gt_ta ta_gt_sa]) (meson ua_in ta_in sa_in Un_iff max.strict_coboundedI1 max.strict_coboundedI2 size_in_args)+ qed have "\f \ ground_heads (head u). extf f (>\<^sub>t) (args u) (args s)" by (clarify, rule extf_trans_from_irrefl[of ?S _ "args t", OF _ _ _ _ _ gt_trans_args]) (auto simp: gt_same_u_t(2,3) gt_same_t_s(3) wary_args wary_u wary_t wary_s gt_irrefl) thus ?thesis by (rule gt_same[OF wt_u_ge_s hd_u_s]) qed qed qed lemma gt_antisym: "wary s \ wary t \ t >\<^sub>t s \ \ s >\<^sub>t t" using gt_irrefl gt_trans by blast subsection \Subterm Property\ lemma gt_sub_fun: "App s t >\<^sub>t s" proof (cases "wt (App s t) >\<^sub>p wt s") case True thus ?thesis using gt_wt by simp next case False hence \_eq_\: "\\<^sub>h = \\<^sub>h" using wt_App_ge_fun dual_order.order_iff_strict wt_App_arg_\\<^sub>h wt_\\<^sub>h_imp_\\<^sub>h_eq_\\<^sub>h unfolding gt_tpoly_def ge_tpoly_def by fast have hd_st: "head (App s t) = head s" by auto have extf: "\f \ ground_heads (head (App s t)). extf f (>\<^sub>t) (args (App s t)) (args s)" by (simp add: \_eq_\ extf_snoc_if_\\<^sub>h_eq_\\<^sub>h) show ?thesis by (rule gt_same[OF wt_App_ge_fun hd_st extf]) qed theorem gt_proper_sub: "wary t \ proper_sub s t \ t >\<^sub>t s" by (induct t) (auto intro: gt_sub_fun gt_sub_arg gt_trans sub.intros wary_sub) subsection \Compatibility with Functions\ lemma gt_compat_fun: assumes wary_t: "wary t" and t'_gt_t: "t' >\<^sub>t t" shows "App s t' >\<^sub>t App s t" proof (rule gt_same; clarify?) show "wt (App s t') \\<^sub>p wt (App s t)" using gt_imp_wt[OF t'_gt_t, unfolded ge_tpoly_def] by (cases s rule: tm_exhaust_apps, auto simp del: apps_append simp: ge_tpoly_def App_apps eval_ztpoly_nonneg intro: ordered_comm_semiring_class.comm_mult_left_mono) next fix f have "extf f (>\<^sub>t) (args s @ [t']) (args s @ [t])" using t'_gt_t by (metis extf_compat_list gt_irrefl[OF wary_t]) thus "extf f (>\<^sub>t) (args (App s t')) (args (App s t))" by simp qed simp theorem gt_compat_fun_strong: assumes wary_t: "wary t" and t'_gt_t: "t' >\<^sub>t t" shows "apps s (t' # us) >\<^sub>t apps s (t # us)" proof (induct us rule: rev_induct) case Nil show ?case using t'_gt_t by (auto intro!: gt_compat_fun[OF wary_t]) next case (snoc u us) note ih = snoc let ?v' = "apps s (t' # us @ [u])" let ?v = "apps s (t # us @ [u])" have "wt ?v' \\<^sub>p wt ?v" using gt_imp_wt[OF ih] by (cases s rule: tm_exhaust_apps, simp del: apps_append add: App_apps apps_append[symmetric] ge_tpoly_def, subst (1 2) zip_eq_butlast_last, simp+) moreover have "head ?v' = head ?v" by simp moreover have "\f \ ground_heads (head ?v'). extf f (>\<^sub>t) (args ?v') (args ?v)" by (metis args_apps extf_compat_list gt_irrefl[OF wary_t] t'_gt_t) ultimately show ?case by (rule gt_same) qed subsection \Compatibility with Arguments\ theorem gt_compat_arg_weak: assumes wary_st: "wary (App s t)" and wary_s't: "wary (App s' t)" and coef_s'_0_ge_s: "coef s' 0 \\<^sub>p coef s 0" and s'_gt_s: "s' >\<^sub>t s" shows "App s' t >\<^sub>t App s t" proof - obtain \ ss where s: "s = apps (Hd \) ss" by (metis tm_exhaust_apps) obtain \' ss' where s': "s' = apps (Hd \') ss'" by (metis tm_exhaust_apps) have len_ss_lt: "of_nat (length ss) < arity_sym\<^sub>h (min_ground_head \)" using wary_st[unfolded s] ground_heads_arity\<^sub>h less_le_trans min_ground_head_in_ground_heads by (metis (no_types) tm_collapse_apps tm_inject_apps wary_AppE\<^sub>h) have \_etc: "\\<^sub>h + \\<^sub>h * (arity_sym\<^sub>h (min_ground_head \) - of_nat (length ss) - 1) = \\<^sub>h * (arity_sym\<^sub>h (min_ground_head \) - of_nat (length ss))" if wary: "wary (App (apps (Hd \) ss) t)" for \ ss proof (cases "\\<^sub>h > 0") case True then obtain n where n: "of_nat n = arity_sym\<^sub>h (min_ground_head \)" by (metis arity_sym\<^sub>h_if_\\<^sub>h_gt_0_E) have "of_nat (length ss) < arity_sym\<^sub>h (min_ground_head \)" using wary by (metis (no_types) wary_AppE\<^sub>h ground_heads_arity\<^sub>h le_less_trans min_ground_head_in_ground_heads not_le tm_collapse_apps tm_inject_apps) thus ?thesis by (fold n, subst of_nat_1[symmetric], fold of_nat_minus_hmset, simp, metis Suc_diff_Suc mult_Suc_right of_nat_add of_nat_mult) qed simp have coef_\'_ge_\: "coef_hd \' (length ss') \\<^sub>p coef_hd \ (length ss)" by (rule coef_s'_0_ge_s[unfolded s s', simplified]) have wt_s'_ge_s: "wt s' \\<^sub>p wt s" by (rule gt_imp_wt[OF s'_gt_s]) have \_tms_len_ss_tms_wt_t_le: "eval_ztpoly A (coef_hd \ (length ss)) * eval_ztpoly A (wt t) \ eval_ztpoly A (coef_hd \' (length ss')) * eval_ztpoly A (wt t)" if legal: "legal_zpassign A" for A using legal coef_\'_ge_\[unfolded ge_tpoly_def] by (simp add: eval_ztpoly_nonneg mult_right_mono) have wt_s't_ge_st: "wt (App s' t) \\<^sub>p wt (App s t)" unfolding s s' by (clarsimp simp del: apps_append simp: App_apps ge_tpoly_def add_ac(1)[symmetric] intro!: add_mono[OF _ \_tms_len_ss_tms_wt_t_le], rule add_le_imp_le_left[of "zhmset_of \\<^sub>h"], unfold add_ac(1)[symmetric] add.commute[of 1] diff_diff_add[symmetric], subst (1 3) ac_simps(3)[unfolded add_ac(1)[symmetric]], subst (1 3) add_ac(1), simp only: zhmset_of_plus[symmetric] \_etc[OF wary_st[unfolded s]] \_etc[OF wary_s't[unfolded s']] add_ac(1) wt_s'_ge_s[unfolded s s', unfolded ge_tpoly_def add_ac(1)[symmetric], simplified]) show ?thesis using s'_gt_s proof cases case gt_wt_s'_s: gt_wt have "wt (App s' t) >\<^sub>p wt (App s t)" unfolding s s' by (clarsimp simp del: apps_append simp: App_apps gt_tpoly_def add_ac(1)[symmetric] intro!: add_less_le_mono[OF _ \_tms_len_ss_tms_wt_t_le], rule add_less_imp_less_left[of "zhmset_of \\<^sub>h"], unfold add_ac(1)[symmetric] add.commute[of 1] diff_diff_add[symmetric], subst (1 3) ac_simps(3)[unfolded add_ac(1)[symmetric]], subst (1 3) add_ac(1), simp only: zhmset_of_plus[symmetric] \_etc[OF wary_st[unfolded s]] \_etc[OF wary_s't[unfolded s']] add_ac(1) gt_wt_s'_s[unfolded s s', unfolded gt_tpoly_def add_ac(1)[symmetric], simplified]) thus ?thesis by (rule gt_wt) next case gt_unary_s'_s: gt_unary have False by (metis ground_heads_arity\<^sub>h gt_unary_s'_s(3) gt_unary_s'_s(4) hmset_of_enat_1 leD of_nat_1 wary_AppE\<^sub>h wary_s't) thus ?thesis by sat next case gt_diff_s'_s: gt_diff show ?thesis by (rule gt_diff[OF wt_s't_ge_st]) (simp add: gt_diff_s'_s(2)) next case gt_same_s'_s: gt_same have hd_s't: "head (App s' t) = head (App s t)" by (simp add: gt_same_s'_s(2)) have "\f \ ground_heads (head (App s' t)). extf f (>\<^sub>t) (args (App s' t)) (args (App s t))" using gt_same_s'_s(3) by (auto intro: extf_compat_append_right) thus ?thesis by (rule gt_same[OF wt_s't_ge_st hd_s't]) qed qed subsection \Stability under Substitution\ primrec subst_zpassign :: "('v \ ('s, 'v) tm) \ ('v pvar \ zhmultiset) \ 'v pvar \ zhmultiset" where "subst_zpassign \ A (PWt x) = eval_ztpoly A (wt (\ x)) - zhmset_of (\\<^sub>h * arity_sym\<^sub>h (min_ground_head (Var x)))" | "subst_zpassign \ A (PCoef x i) = eval_ztpoly A (coef (\ x) i)" lemma legal_subst_zpassign: assumes legal: "legal_zpassign A" and wary_\: "wary_subst \" shows "legal_zpassign (subst_zpassign \ A)" unfolding legal_zpassign_def proof fix v show "subst_zpassign \ A v \ min_zpassign v" proof (cases v) case v: (PWt x) obtain \ ss where \x: "\ x = apps (Hd \) ss" by (rule tm_exhaust_apps) have ghd_\: "ground_heads \ \ ground_heads_var x" using wary_\[unfolded wary_subst_def, rule_format, of x, unfolded \x] by simp have "zhmset_of (wt_sym (min_ground_head (Var x)) + \\<^sub>h * arity_sym\<^sub>h (min_ground_head (Var x))) \ eval_ztpoly A (wt0 \) + zhmset_of (\\<^sub>h * arity_sym\<^sub>h (min_ground_head \))" proof - have mgh_x_min: "zhmset_of (wt_sym (min_ground_head (Var x)) + \\<^sub>h * arity_sym\<^sub>h (min_ground_head (Var x))) \ zhmset_of (wt_sym (min_ground_head \) + \\<^sub>h * arity_sym\<^sub>h (min_ground_head \))" by (simp add: zmset_of_le zhmset_of_le ghd_\ min_ground_head_antimono) have wt_mgh_le_wt0: "zhmset_of (wt_sym (min_ground_head \)) \ eval_ztpoly A (wt0 \)" using wt0_ge_min_ground_head[OF legal] by blast show ?thesis by (rule order_trans[OF mgh_x_min]) (simp add: zhmset_of_plus wt_mgh_le_wt0) qed also have "\ \ eval_ztpoly A (wt0 \) + zhmset_of ((\\<^sub>h * (arity_sym\<^sub>h (min_ground_head \) - of_nat (length ss))) + of_nat (length ss) * \\<^sub>h)" proof - have "zhmset_of (\\<^sub>h * arity_sym\<^sub>h (min_ground_head \)) \ zhmset_of (\\<^sub>h * (of_nat (length ss) + (arity_sym\<^sub>h (min_ground_head \) - of_nat (length ss))))" by (metis add.commute le_minus_plus_same_hmset mult_le_mono2_hmset zhmset_of_le) thus ?thesis by (simp add: add.commute add.left_commute distrib_left mult.commute) qed also have "\ \ eval_ztpoly A (wt0 \) + zhmset_of ((\\<^sub>h * (arity_sym\<^sub>h (min_ground_head \) - of_nat (length ss))) + of_nat (length ss) * \\<^sub>h)" using \\<^sub>h_le_\\<^sub>h zhmset_of_le by auto also have "\ \ eval_ztpoly A (wt0 \) + zhmset_of (\\<^sub>h * (arity_sym\<^sub>h (min_ground_head \) - of_nat (length ss))) + wt_args 0 A \ ss" using wt_args_ge_length_times_\\<^sub>h[OF legal] by (simp add: algebra_simps zhmset_of_plus zhmset_of_times of_nat_zhmset) finally have wt_x_le_\ssts: "zhmset_of (wt_sym (min_ground_head (Var x)) + \\<^sub>h * arity_sym\<^sub>h (min_ground_head (Var x))) \ eval_ztpoly A (wt0 \) + zhmset_of (\\<^sub>h * (arity_sym\<^sub>h (min_ground_head \) - of_nat (length ss))) + wt_args 0 A \ ss" by assumption show ?thesis using wt_x_le_\ssts[unfolded wt_args_def] by (simp add: v \x comp_def le_diff_eq add.assoc[symmetric] ZHMSet_plus[symmetric] zmset_of_plus[symmetric] hmsetmset_plus[symmetric] zmset_of_le) next case (PCoef x i) thus ?thesis using coef_gt_0[OF legal, unfolded zero_less_iff_1_le_hmset] by (simp add: zhmset_of_1 zero_less_iff_1_le_zhmset) qed qed lemma wt_subst: assumes legal: "legal_zpassign A" and wary_\: "wary_subst \" shows "wary s \ eval_ztpoly A (wt (subst \ s)) = eval_ztpoly (subst_zpassign \ A) (wt s)" proof (induct s rule: tm_induct_apps) case (apps \ ss) note ih = this(1) and wary_\ss = this(2) have wary_nth_ss: "\i. i < length ss \ wary (ss ! i)" using wary_args[OF _ wary_\ss] by force show ?case proof (cases \) case \: (Var x) show ?thesis proof (cases "\ x" rule: tm_exhaust_apps) case \x: (apps \ ts) have wary_\x: "wary (\ x)" using wary_\ wary_subst_def by blast have coef_subst: "\i. eval_tpoly A (zhmset_of_tpoly (coef_hd \ (i + length ts))) = eval_tpoly (subst_zpassign \ A) (zhmset_of_tpoly (coef_hd (Var x) i))" by (simp add: \x) have tedious_ary_arith: "arity_sym\<^sub>h (min_ground_head (Var x)) + (arity_sym\<^sub>h (min_ground_head \) - (of_nat (length ss) + of_nat (length ts))) = arity_sym\<^sub>h (min_ground_head \) - of_nat (length ts) + (arity_sym\<^sub>h (min_ground_head (Var x)) - of_nat (length ss))" if \_gt_0: "\\<^sub>h > 0" proof - obtain m where m: "of_nat m = arity_sym\<^sub>h (min_ground_head (Var x))" by (metis arity_sym\<^sub>h_if_\\<^sub>h_gt_0_E[OF \_gt_0]) obtain n where n: "of_nat n = arity_sym\<^sub>h (min_ground_head \)" by (metis arity_sym\<^sub>h_if_\\<^sub>h_gt_0_E[OF \_gt_0]) have "m \ length ss" unfolding of_nat_le_hmset[symmetric] m using wary_\ss[unfolded \] by (cases rule: wary_cases_apps\<^sub>h, clarsimp, metis arity_hd.simps(1) enat_ile enat_ord_simps(1) ground_heads_arity hmset_of_enat_inject hmset_of_enat_of_nat le_trans m min_ground_head_in_ground_heads of_nat_eq_enat of_nat_le_hmset_of_enat_iff) moreover have n_ge_len_ss_ts: "n \ length ss + length ts" proof - have "of_nat (length ss) + of_nat (length ts) \ arity_hd\<^sub>h \ + of_nat (length ts)" using wary_\ss wary_cases_apps\<^sub>h by fastforce also have "\ = arity_var\<^sub>h x + of_nat (length ts)" by (simp add: \) also have "\ \ arity\<^sub>h (\ x) + of_nat (length ts)" using wary_\ wary_subst_def by auto also have "\ = arity\<^sub>h (apps (Hd \) ts) + of_nat (length ts)" by (simp add: \x) also have "\ = arity_hd\<^sub>h \" using wary_\x[unfolded \x] by (cases rule: wary_cases_apps\<^sub>h, cases "arity_hd \", simp add: of_nat_add[symmetric] of_nat_minus_hmset[symmetric], metis \_gt_0 arity_hd_ne_infinity_if_\_gt_0 of_nat_0 of_nat_less_hmset) also have "\ \ arity_sym\<^sub>h (min_ground_head \)" using ground_heads_arity\<^sub>h min_ground_head_in_ground_heads by blast finally show ?thesis unfolding of_nat_le_hmset[symmetric] n by simp qed moreover have "n \ length ts" using n_ge_len_ss_ts by simp ultimately show ?thesis by (fold m n of_nat_add of_nat_minus_hmset, unfold of_nat_inject_hmset, fastforce) qed have "eval_tpoly A (zhmset_of_tpoly (wt (subst \ (apps (Hd (Var x)) ss)))) = eval_tpoly A (zhmset_of_tpoly (wt0 \)) + zhmset_of (\\<^sub>h * (arity_sym\<^sub>h (min_ground_head \) - (of_nat (length ts) + of_nat (length ss)))) + wt_args 0 A \ (ts @ map (subst \) ss)" by (simp del: apps_append add: apps_append[symmetric] \x wt_args_def comp_def) also have "\ = eval_tpoly A (zhmset_of_tpoly (wt0 \)) + zhmset_of (\\<^sub>h * (arity_sym\<^sub>h (min_ground_head \) - (of_nat (length ts) + of_nat (length ss)))) + wt_args 0 A \ ts + wt_args (length ts) A \ (map (subst \) ss)" by (simp add: wt_args_def zip_append_0_upt[of ts "map (subst \) ss", simplified]) also have "\ = eval_tpoly A (zhmset_of_tpoly (wt0 \)) + zhmset_of (\\<^sub>h * (arity_sym\<^sub>h (min_ground_head \) - (of_nat (length ts) + of_nat (length ss)))) + wt_args 0 A \ ts + wt_args 0 (subst_zpassign \ A) (Var x) ss" by (auto intro!: arg_cong[of _ _ sum_list] nth_map_conv simp: wt_args_def coef_subst add.commute zhmset_of_times ih[OF nth_mem wary_nth_ss]) also have "\ = eval_tpoly (subst_zpassign \ A) (zhmset_of_tpoly (wt0 (Var x))) + zhmset_of (\\<^sub>h * (arity_sym\<^sub>h (min_ground_head (Var x)) - of_nat (length ss))) + wt_args 0 (subst_zpassign \ A) (Var x) ss" by (simp add: \x wt_args_def comp_def algebra_simps ring_distribs(1)[symmetric] zhmset_of_times zhmset_of_plus[symmetric] zhmset_of_0[symmetric]) (use tedious_ary_arith in fastforce) also have "\ = eval_tpoly (subst_zpassign \ A) (zhmset_of_tpoly (wt (apps (Hd (Var x)) ss)))" by (simp add: wt_args_def comp_def) finally show ?thesis unfolding \ by assumption qed next case \: (Sym f) have "eval_tpoly A (zhmset_of_tpoly (wt (subst \ (apps (Hd (Sym f)) ss)))) = zhmset_of (wt_sym f) + zhmset_of (\\<^sub>h * (arity_sym\<^sub>h f - of_nat (length ss))) + wt_args 0 A (Sym f) (map (subst \) ss)" by (simp add: wt_args_def comp_def) also have "\ = zhmset_of (wt_sym f) + zhmset_of (\\<^sub>h * (arity_sym\<^sub>h f - of_nat (length ss))) + wt_args 0 (subst_zpassign \ A) (Sym f) ss" by (auto simp: wt_args_def ih[OF _ wary_nth_ss] intro!: arg_cong[of _ _ sum_list] nth_map_conv) also have "\ = eval_tpoly (subst_zpassign \ A) (zhmset_of_tpoly (wt (apps (Hd (Sym f)) ss)))" by (simp add: wt_args_def comp_def) finally show ?thesis unfolding \ by assumption qed qed theorem gt_subst: assumes wary_\: "wary_subst \" shows "wary t \ wary s \ t >\<^sub>t s \ subst \ t >\<^sub>t subst \ s" proof (simp only: atomize_imp, rule measure_induct_rule[of "\(t, s). {#size t, size s#}" "\(t, s). wary t \ wary s \ t >\<^sub>t s \ subst \ t >\<^sub>t subst \ s" "(t, s)", simplified prod.case], simp only: split_paired_all prod.case atomize_imp[symmetric]) fix t s assume ih: "\ta sa. {#size ta, size sa#} < {#size t, size s#} \ wary ta \ wary sa \ ta >\<^sub>t sa \ subst \ ta >\<^sub>t subst \ sa" and wary_t: "wary t" and wary_s: "wary s" and t_gt_s: "t >\<^sub>t s" show "subst \ t >\<^sub>t subst \ s" using t_gt_s proof cases case gt_wt_t_s: gt_wt have "wt (subst \ t) >\<^sub>p wt (subst \ s)" by (auto simp: gt_tpoly_def wary_s wary_t wt_subst[OF _ wary_\] intro: gt_wt_t_s[unfolded gt_tpoly_def, rule_format] elim: legal_subst_zpassign[OF _ wary_\]) thus ?thesis by (rule gt_wt) next assume wt_t_ge_s: "wt t \\<^sub>p wt s" have wt_\t_ge_\s: "wt (subst \ t) \\<^sub>p wt (subst \ s)" by (auto simp: ge_tpoly_def wary_s wary_t wt_subst[OF _ wary_\] intro: wt_t_ge_s[unfolded ge_tpoly_def, rule_format] elim: legal_subst_zpassign[OF _ wary_\]) { case gt_unary have wary_\t: "wary (subst \ t)" by (simp add: wary_subst_wary wary_t wary_\) show ?thesis proof (cases t) case Hd hence False using gt_unary(3) by simp thus ?thesis by sat next case t: (App t1 t2) hence t2: "t2 = arg t" by simp hence wary_t2: "wary t2" using wary_t by blast show ?thesis proof (cases "t2 = s") case True moreover have "subst \ t >\<^sub>t subst \ t2" using gt_sub_arg wary_\t unfolding t by simp ultimately show ?thesis by simp next case t2_ne_s: False hence t2_gt_s: "t2 >\<^sub>t s" using gt_unary(5) t2 by blast have "subst \ t2 >\<^sub>t subst \ s" by (rule ih[OF _ wary_t2 wary_s t2_gt_s]) (simp add: t) thus ?thesis by (metis gt_sub_arg gt_trans subst.simps(2) t wary_\ wary_\t wary_s wary_subst_wary wary_t2) qed qed } { case _: gt_diff note hd_t_gt_hd_s = this(2) have "head (subst \ t) >\<^sub>h\<^sub>d head (subst \ s)" by (meson hd_t_gt_hd_s wary_subst_ground_heads gt_hd_def rev_subsetD wary_\) thus ?thesis by (rule gt_diff[OF wt_\t_ge_\s]) } { case _: gt_same note hd_s_eq_hd_t = this(2) and extf = this(3) have hd_\t: "head (subst \ t) = head (subst \ s)" by (simp add: hd_s_eq_hd_t) { fix f assume f_in_grs: "f \ ground_heads (head (subst \ t))" let ?S = "set (args t) \ set (args s)" have extf_args_s_t: "extf f (>\<^sub>t) (args t) (args s)" using extf f_in_grs wary_subst_ground_heads wary_\ by blast have "extf f (>\<^sub>t) (map (subst \) (args t)) (map (subst \) (args s))" proof (rule extf_map[of ?S, OF _ _ _ _ _ _ extf_args_s_t]) show "\x \ ?S. \ subst \ x >\<^sub>t subst \ x" using gt_irrefl wary_t wary_s wary_args wary_\ wary_subst_wary by fastforce next show "\z \ ?S. \y \ ?S. \x \ ?S. subst \ z >\<^sub>t subst \ y \ subst \ y >\<^sub>t subst \ x \ subst \ z >\<^sub>t subst \ x" using gt_trans wary_t wary_s wary_args wary_\ wary_subst_wary by (metis Un_iff) next have sz_a: "\ta \ ?S. \sa \ ?S. {#size ta, size sa#} < {#size t, size s#}" by (fastforce intro: Max_lt_imp_lt_mset dest: size_in_args) show "\y \ ?S. \x \ ?S. y >\<^sub>t x \ subst \ y >\<^sub>t subst \ x" using ih sz_a size_in_args wary_t wary_s wary_args wary_\ wary_subst_wary by fastforce qed auto hence "extf f (>\<^sub>t) (args (subst \ t)) (args (subst \ s))" by (auto simp: hd_s_eq_hd_t intro: extf_compat_append_left) } hence "\f \ ground_heads (head (subst \ t)). extf f (>\<^sub>t) (args (subst \ t)) (args (subst \ s))" by blast thus ?thesis by (rule gt_same[OF wt_\t_ge_\s hd_\t]) } qed qed subsection \Totality on Ground Terms\ lemma wt_total_ground: assumes gr_t: "ground t" and gr_s: "ground s" shows "wt t >\<^sub>p wt s \ wt s >\<^sub>p wt t \ wt t =\<^sub>p wt s" unfolding gt_tpoly_def eq_tpoly_def by (subst (1 2 3) ground_eval_ztpoly_wt_eq[OF gr_t, of _ undefined], subst (1 2 3) ground_eval_ztpoly_wt_eq[OF gr_s, of _ undefined], auto) theorem gt_total_ground: assumes extf_total: "\f. ext_total (extf f)" shows "ground t \ ground s \ t >\<^sub>t s \ s >\<^sub>t t \ t = s" proof (simp only: atomize_imp, rule measure_induct_rule[of "\(t, s). {# size t, size s #}" "\(t, s). ground t \ ground s \ t >\<^sub>t s \ s >\<^sub>t t \ t = s" "(t, s)", simplified prod.case], simp only: split_paired_all prod.case atomize_imp[symmetric]) fix t s :: "('s, 'v) tm" assume ih: "\ta sa. {# size ta, size sa #} < {# size t, size s #} \ ground ta \ ground sa \ ta >\<^sub>t sa \ sa >\<^sub>t ta \ ta = sa" and gr_t: "ground t" and gr_s: "ground s" let ?case = "t >\<^sub>t s \ s >\<^sub>t t \ t = s" { assume "wt t >\<^sub>p wt s" hence "t >\<^sub>t s" by (rule gt_wt) } moreover { assume "wt s >\<^sub>p wt t" hence "s >\<^sub>t t" by (rule gt_wt) } moreover { assume "wt t =\<^sub>p wt s" hence wt_t_ge_s: "wt t \\<^sub>p wt s" and wt_s_ge_t: "wt s \\<^sub>p wt t" by (simp add: eq_tpoly_def ge_tpoly_def)+ obtain g where \: "head t = Sym g" by (metis ground_head[OF gr_t] hd.collapse(2)) obtain f where \: "head s = Sym f" by (metis ground_head[OF gr_s] hd.collapse(2)) { assume g_gt_f: "g >\<^sub>s f" have "t >\<^sub>t s" by (rule gt_diff[OF wt_t_ge_s]) (simp add: \ \ g_gt_f gt_hd_def) } moreover { assume f_gt_g: "f >\<^sub>s g" have "s >\<^sub>t t" by (rule gt_diff[OF wt_s_ge_t]) (simp add: \ \ f_gt_g gt_hd_def) } moreover { assume g_eq_f: "g = f" hence hd_t: "head t = head s" using \ \ by force note hd_s = hd_t[symmetric] let ?ts = "args t" let ?ss = "args s" have gr_ts: "\t \ set ?ts. ground t" using gr_t ground_args by auto have gr_ss: "\s \ set ?ss. ground s" using gr_s ground_args by auto have ?case proof (cases "?ts = ?ss") case ts_eq_ss: True show ?thesis using \ \ g_eq_f ts_eq_ss by (simp add: tm_expand_apps) next case False hence "extf g (>\<^sub>t) ?ts ?ss \ extf g (>\<^sub>t) ?ss ?ts" using ih gr_ss gr_ts less_multiset_doubletons ext_total.total[OF extf_total, rule_format, of "set ?ts \ set ?ss" "(>\<^sub>t)" ?ts ?ss g] by (metis Un_commute Un_iff in_lists_iff_set size_in_args sup_ge2) moreover { assume extf: "extf g (>\<^sub>t) ?ts ?ss" have "t >\<^sub>t s" by (rule gt_same[OF wt_t_ge_s hd_t]) (simp add: extf \) } moreover { assume extf: "extf g (>\<^sub>t) ?ss ?ts" have "s >\<^sub>t t" by (rule gt_same[OF wt_s_ge_t hd_s]) (simp add: extf[unfolded g_eq_f] \) } ultimately show ?thesis by sat qed } ultimately have ?case using gt_sym_total by blast } ultimately show ?case using wt_total_ground[OF gr_t gr_s] by fast qed subsection \Well-foundedness\ abbreviation gtw :: "('s, 'v) tm \ ('s, 'v) tm \ bool" (infix ">\<^sub>t\<^sub>w" 50) where "(>\<^sub>t\<^sub>w) \ \t s. wary t \ wary s \ t >\<^sub>t s" abbreviation gtwg :: "('s, 'v) tm \ ('s, 'v) tm \ bool" (infix ">\<^sub>t\<^sub>w\<^sub>g" 50) where "(>\<^sub>t\<^sub>w\<^sub>g) \ \t s. ground t \ t >\<^sub>t\<^sub>w s" lemma ground_gt_unary: assumes gr_t: "ground t" shows "\ gt_unary t s" proof assume gt_unary_t_s: "gt_unary t s" hence "t >\<^sub>t s" using gt_iff_wt_unary_diff_same by blast hence gr_s: "ground s" using gr_t gt_imp_vars by blast have ngr_t_or_s: "\ ground t \ \ ground s" using gt_unary_t_s by cases (blast dest: ground_head not_comp_hd_imp_Var) show False using gr_t gr_s ngr_t_or_s by sat qed theorem gt_wf: "wfP (\s t. t >\<^sub>t\<^sub>w s)" proof - have ground_wfP: "wfP (\s t. t >\<^sub>t\<^sub>w\<^sub>g s)" unfolding wfP_iff_no_inf_chain proof assume "\f. inf_chain (>\<^sub>t\<^sub>w\<^sub>g) f" then obtain t where t_bad: "bad (>\<^sub>t\<^sub>w\<^sub>g) t" unfolding inf_chain_def bad_def by blast let ?ff = "worst_chain (>\<^sub>t\<^sub>w\<^sub>g) (\t s. size t > size s)" let ?A = min_passign note wf_sz = wf_app[OF wellorder_class.wf, of size, simplified] have ffi_ground: "\i. ground (?ff i)" and ffi_wary: "\i. wary (?ff i)" using worst_chain_bad[OF wf_sz t_bad, unfolded inf_chain_def] by fast+ have "inf_chain (>\<^sub>t\<^sub>w\<^sub>g) ?ff" by (rule worst_chain_bad[OF wf_sz t_bad]) hence bad_wt_diff_same: "inf_chain (\t s. ground t \ (gt_wt t s \ gt_diff t s \ gt_same t s)) ?ff" unfolding inf_chain_def using gt_iff_wt_unary_diff_same ground_gt_unary by blast have wf_wt: "wf {(s, t). ground t \ gt_wt t s}" by (rule wf_subset[OF wf_app[of _ "eval_tpoly ?A \ wt", OF wf_less_hmultiset]], simp add: gt_wt.simps gt_tpoly_def, fold zhmset_of_less, auto simp: legal_min_zpassign gt_wt.simps gt_tpoly_def) have wt_O_diff_same: "{(s, t). ground t \ gt_wt t s} O {(s, t). ground t \ wt t =\<^sub>p wt s \ (gt_diff t s \ gt_same t s)} \ {(s, t). ground t \ gt_wt t s}" unfolding gt_wt.simps gt_diff.simps gt_same.simps by (auto intro: ge_gt_tpoly_trans) have wt_diff_same_as_union: "{(s, t). ground t \ (gt_wt t s \ gt_diff t s \ gt_same t s)} = {(s, t). ground t \ gt_wt t s} \ {(s, t). ground t \ wt t =\<^sub>p wt s \ (gt_diff t s \ gt_same t s)}" using gt_ge_tpoly_trans gt_tpoly_irrefl wt_ge_vars wt_total_ground by (fastforce simp: gt_wt.simps gt_diff.simps gt_same.simps) obtain k1 where bad_diff_same: "inf_chain (\t s. ground t \ wt t =\<^sub>p wt s \ (gt_diff t s \ gt_same t s)) (\i. ?ff (i + k1))" using wf_infinite_down_chain_compatible[OF wf_wt _ wt_O_diff_same, of ?ff] bad_wt_diff_same unfolding inf_chain_def wt_diff_same_as_union[symmetric] by auto have "wf {(s, t). ground s \ ground t \ wt t =\<^sub>p wt s \ sym (head t) >\<^sub>s sym (head s)}" using gt_sym_wf unfolding wfP_def wf_iff_no_infinite_down_chain by fast moreover have "{(s, t). ground t \ wt t =\<^sub>p wt s \ gt_diff t s} \ {(s, t). ground s \ ground t \ wt t =\<^sub>p wt s \ sym (head t) >\<^sub>s sym (head s)}" proof (clarsimp, intro conjI) fix s t assume gr_t: "ground t" and gt_diff_t_s: "gt_diff t s" thus gr_s: "ground s" using gt_iff_wt_unary_diff_same gt_imp_vars by fastforce show "sym (head t) >\<^sub>s sym (head s)" using gt_diff_t_s by cases (simp add: gt_hd_def gr_s gr_t ground_hd_in_ground_heads) qed ultimately have wf_diff: "wf {(s, t). ground t \ wt t =\<^sub>p wt s \ gt_diff t s}" by (rule wf_subset) have diff_O_same: "{(s, t). ground t \ wt t =\<^sub>p wt s \ gt_diff t s} O {(s, t). ground t \ wt t =\<^sub>p wt s \ gt_same t s} \ {(s, t). ground t \ wt t =\<^sub>p wt s \ gt_diff t s}" unfolding gt_diff.simps gt_same.simps by (auto intro: ge_ge_tpoly_trans simp: eq_tpoly_def) have diff_same_as_union: "{(s, t). ground t \ wt t =\<^sub>p wt s \ (gt_diff t s \ gt_same t s)} = {(s, t). ground t \ wt t =\<^sub>p wt s \ gt_diff t s} \ {(s, t). ground t \ wt t =\<^sub>p wt s \ gt_same t s}" by auto obtain k2 where bad_same: "inf_chain (\t s. ground t \ wt t =\<^sub>p wt s \ gt_same t s) (\i. ?ff (i + k2))" using wf_infinite_down_chain_compatible[OF wf_diff _ diff_O_same, of "\i. ?ff (i + k1)"] bad_diff_same unfolding inf_chain_def diff_same_as_union[symmetric] by (auto simp: add.assoc) hence hd_sym: "\i. is_Sym (head (?ff (i + k2)))" unfolding inf_chain_def by (simp add: ground_head) define f where "f = sym (head (?ff k2))" define w where "w = eval_tpoly ?A (wt (?ff k2))" have "head (?ff (i + k2)) = Sym f \ eval_tpoly ?A (wt (?ff (i + k2))) = w" for i proof (induct i) case 0 thus ?case by (auto simp: f_def w_def hd.collapse(2)[OF hd_sym, of 0, simplified]) next case (Suc ia) thus ?case using bad_same unfolding inf_chain_def gt_same.simps zhmset_of_inject[symmetric] by (simp add: eq_tpoly_def legal_min_zpassign) qed note hd_eq_f = this[THEN conjunct1] and wt_eq_w = this[THEN conjunct2] define max_args where "max_args = (if \\<^sub>h = 0 then sum_coefs w else the_enat (arity_sym f))" have nargs_le_max_args: "num_args (?ff (i + k2)) \ max_args" for i proof (cases "\\<^sub>h = 0") case \_ne_0: False hence ary_f_ne_inf: "arity_sym f \ \" using arity_sym_ne_infinity_if_\_gt_0 of_nat_0 by blast have "enat (num_args (worst_chain (\t s. ground t \ t >\<^sub>t\<^sub>w s) (\t s. size s < size t) (i + k2))) \ arity_sym f" using wary_num_args_le_arity_head[OF ffi_wary[of "i + k2"]] by (simp add: hd_eq_f) with \_ne_0 show ?thesis by (simp del: enat_ord_simps add: max_args_def enat_ord_simps(1)[symmetric] enat_the_enat_iden[OF ary_f_ne_inf]) next case \_eq_0: True show ?thesis using sum_coefs_ge_num_args_if_\\<^sub>h_eq_0[OF legal_min_passign \_eq_0 ffi_wary[of "i + k2"]] by (simp add: max_args_def \_eq_0 wt_eq_w) qed let ?U_of = "\i. set (args (?ff (i + k2)))" define U where "U = (\i. ?U_of i)" have gr_u: "\u. u \ U \ ground u" unfolding U_def by (blast dest: ground_args[OF _ ffi_ground]) have wary_u: "\u. u \ U \ wary u" unfolding U_def by (blast dest: wary_args[OF _ ffi_wary]) have "\ bad (>\<^sub>t\<^sub>w\<^sub>g) u" if u_in: "u \ ?U_of i" for u i proof assume u_bad: "bad (>\<^sub>t\<^sub>w\<^sub>g) u" have sz_u: "size u < size (?ff (i + k2))" by (rule size_in_args[OF u_in]) show False proof (cases "i + k2") case 0 thus False using sz_u min_worst_chain_0[OF wf_sz u_bad] by simp next case Suc hence gt: "?ff (i + k2 - 1) >\<^sub>t\<^sub>w ?ff (i + k2)" using worst_chain_pred[OF wf_sz t_bad] by auto moreover have "?ff (i + k2) >\<^sub>t\<^sub>w u" using gt gt_proper_sub sub_args sz_u u_in wary_args by auto ultimately have "?ff (i + k2 - 1) >\<^sub>t\<^sub>w u" using gt_trans by blast thus False using Suc sz_u min_worst_chain_Suc[OF wf_sz u_bad] ffi_ground by fastforce qed qed hence u_good: "\u. u \ U \ \ bad (>\<^sub>t\<^sub>w\<^sub>g) u" unfolding U_def by blast let ?gtwu = "\t s. t \ U \ t >\<^sub>t\<^sub>w s" have gtwu_irrefl: "\x. \ ?gtwu x x" using gt_irrefl by auto have "\i j. \t \ set (args (?ff (i + k2))). \s \ set (args (?ff (j + k2))). t >\<^sub>t s \ t \ U \ t >\<^sub>t\<^sub>w s" using wary_u unfolding U_def by blast moreover have "\i. extf f (>\<^sub>t) (args (?ff (i + k2))) (args (?ff (Suc i + k2)))" using bad_same hd_eq_f unfolding inf_chain_def gt_same.simps by auto ultimately have "\i. extf f ?gtwu (args (?ff (i + k2))) (args (?ff (Suc i + k2)))" by (rule extf_mono_strong) hence "inf_chain (extf f ?gtwu) (\i. args (?ff (i + k2)))" unfolding inf_chain_def by blast hence nwf_ext: "\ wfP (\xs ys. length ys \ max_args \ length xs \ max_args \ extf f ?gtwu ys xs)" unfolding inf_chain_def wfP_def wf_iff_no_infinite_down_chain using nargs_le_max_args by fast have gtwu_le_gtwg: "?gtwu \ (>\<^sub>t\<^sub>w\<^sub>g)" by (auto intro!: gr_u) have "wfP (\s t. ?gtwu t s)" unfolding wfP_iff_no_inf_chain proof (intro notI, elim exE) fix f assume bad_f: "inf_chain ?gtwu f" hence bad_f0: "bad ?gtwu (f 0)" by (rule inf_chain_bad) hence "f 0 \ U" using bad_f unfolding inf_chain_def by blast hence "\ bad (>\<^sub>t\<^sub>w\<^sub>g) (f 0)" using u_good by blast hence "\ bad ?gtwu (f 0)" using bad_f inf_chain_bad inf_chain_subset[OF _ gtwu_le_gtwg] by blast thus False using bad_f0 by sat qed hence wf_ext: "wfP (\xs ys. length ys \ max_args \ length xs \ max_args \ extf f ?gtwu ys xs)" using extf_wf_bounded[of ?gtwu] gtwu_irrefl by blast show False using nwf_ext wf_ext by blast qed let ?subst = "subst grounding_\" have "wfP (\s t. ?subst t >\<^sub>t\<^sub>w\<^sub>g ?subst s)" by (rule wfP_app[OF ground_wfP]) hence "wfP (\s t. ?subst t >\<^sub>t\<^sub>w ?subst s)" by (simp add: ground_grounding_\) thus ?thesis by (auto intro: wfP_subset wary_subst_wary[OF wary_grounding_\] gt_subst[OF wary_grounding_\]) qed end end diff --git a/thys/Lambda_Free_RPOs/Lambda_Free_Util.thy b/thys/Lambda_Free_RPOs/Lambda_Free_Util.thy --- a/thys/Lambda_Free_RPOs/Lambda_Free_Util.thy +++ b/thys/Lambda_Free_RPOs/Lambda_Free_Util.thy @@ -1,629 +1,626 @@ (* Title: Utilities for Lambda-Free Orders Author: Jasmin Blanchette , 2016 Maintainer: Jasmin Blanchette *) section \Utilities for Lambda-Free Orders\ theory Lambda_Free_Util imports "HOL-Library.Extended_Nat" "HOL-Library.Multiset_Order" begin text \ This theory gathers various lemmas that likely belong elsewhere in Isabelle or the \emph{Archive of Formal Proofs}. Most (but certainly not all) of them are used to formalize orders on \\\-free higher-order terms. \ -hide_const (open) Complex.arg - - subsection \Finite Sets\ lemma finite_set_fold_singleton[simp]: "Finite_Set.fold f z {x} = f x z" proof - have "fold_graph f z {x} (f x z)" by (auto intro: fold_graph.intros) moreover { fix X y have "fold_graph f z X y \ (X = {} \ y = z) \ (X = {x} \ y = f x z)" by (induct rule: fold_graph.induct) auto } ultimately have "(THE y. fold_graph f z {x} y) = f x z" by blast thus ?thesis by (simp add: Finite_Set.fold_def) qed subsection \Function Power\ lemma funpow_lesseq_iter: fixes f :: "('a::order) \ 'a" assumes mono: "\k. k \ f k" and m_le_n: "m \ n" shows "(f ^^ m) k \ (f ^^ n) k" using m_le_n by (induct n) (fastforce simp: le_Suc_eq intro: mono order_trans)+ lemma funpow_less_iter: fixes f :: "('a::order) \ 'a" assumes mono: "\k. k < f k" and m_lt_n: "m < n" shows "(f ^^ m) k < (f ^^ n) k" using m_lt_n by (induct n) (auto, blast intro: mono less_trans dest: less_antisym) subsection \Least Operator\ lemma Least_eq[simp]: "(LEAST y. y = x) = x" and "(LEAST y. x = y) = x" for x :: "'a::order" by (blast intro: Least_equality)+ lemma Least_in_nonempty_set_imp_ex: fixes f :: "'b \ ('a::wellorder)" assumes A_nemp: "A \ {}" and P_least: "P (LEAST y. \x \ A. y = f x)" shows "\x \ A. P (f x)" proof - obtain a where a: "a \ A" using A_nemp by fast have "\x. x \ A \ (LEAST y. \x. x \ A \ y = f x) = f x" by (rule LeastI[of _ "f a"]) (fast intro: a) thus ?thesis by (metis P_least) qed lemma Least_eq_0_enat: "P 0 \ (LEAST x :: enat. P x) = 0" by (simp add: Least_equality) subsection \Antisymmetric Relations\ lemma irrefl_trans_imp_antisym: "irrefl r \ trans r \ antisym r" unfolding irrefl_def trans_def antisym_def by fast lemma irreflp_transp_imp_antisymP: "irreflp p \ transp p \ antisymp p" by (fact irrefl_trans_imp_antisym [to_pred]) subsection \Acyclic Relations\ lemma finite_nonempty_ex_succ_imp_cyclic: assumes fin: "finite A" and nemp: "A \ {}" and ex_y: "\x \ A. \y \ A. (y, x) \ r" shows "\ acyclic r" proof - let ?R = "{(x, y). x \ A \ y \ A \ (x, y) \ r}" have R_sub_r: "?R \ r" by auto have "?R \ A \ A" by auto hence fin_R: "finite ?R" by (auto intro: fin dest!: infinite_super) have "\ acyclic ?R" by (rule notI, drule finite_acyclic_wf[OF fin_R], unfold wf_eq_minimal, drule spec[of _ A], use ex_y nemp in blast) thus ?thesis using R_sub_r acyclic_subset by auto qed subsection \Reflexive, Transitive Closure\ lemma relcomp_subset_left_imp_relcomp_trancl_subset_left: assumes sub: "R O S \ R" shows "R O S\<^sup>* \ R" proof fix x assume "x \ R O S\<^sup>*" then obtain n where "x \ R O S ^^ n" using rtrancl_imp_relpow by fastforce thus "x \ R" proof (induct n) case (Suc m) thus ?case by (metis (no_types) O_assoc inf_sup_ord(3) le_iff_sup relcomp_distrib2 relpow.simps(2) relpow_commute sub subsetCE) qed auto qed lemma f_chain_in_rtrancl: assumes m_le_n: "m \ n" and f_chain: "\i \ {m.. R" shows "(f m, f n) \ R\<^sup>*" proof (rule relpow_imp_rtrancl, rule relpow_fun_conv[THEN iffD2], intro exI conjI) let ?g = "\i. f (m + i)" let ?k = "n - m" show "?g 0 = f m" by simp show "?g ?k = f n" using m_le_n by force show "(\i < ?k. (?g i, ?g (Suc i)) \ R)" by (simp add: f_chain) qed lemma f_rev_chain_in_rtrancl: assumes m_le_n: "m \ n" and f_chain: "\i \ {m.. R" shows "(f n, f m) \ R\<^sup>*" by (rule f_chain_in_rtrancl[OF m_le_n, of "\i. f (n + m - i)", simplified]) (metis f_chain le_add_diff Suc_diff_Suc Suc_leI atLeastLessThan_iff diff_Suc_diff_eq1 diff_less le_add1 less_le_trans zero_less_Suc) subsection \Well-Founded Relations\ lemma wf_app: "wf r \ wf {(x, y). (f x, f y) \ r}" unfolding wf_eq_minimal by (intro allI, drule spec[of _ "f ` Q" for Q]) fast lemma wfP_app: "wfP p \ wfP (\x y. p (f x) (f y))" unfolding wfP_def by (rule wf_app[of "{(x, y). p x y}" f, simplified]) lemma wf_exists_minimal: assumes wf: "wf r" and Q: "Q x" shows "\x. Q x \ (\y. (f y, f x) \ r \ \ Q y)" using wf_eq_minimal[THEN iffD1, OF wf_app[OF wf], rule_format, of _ "{x. Q x}", simplified, OF Q] by blast lemma wfP_exists_minimal: assumes wf: "wfP p" and Q: "Q x" shows "\x. Q x \ (\y. p (f y) (f x) \ \ Q y)" by (rule wf_exists_minimal[of "{(x, y). p x y}" Q x, OF wf[unfolded wfP_def] Q, simplified]) lemma finite_irrefl_trans_imp_wf: "finite r \ irrefl r \ trans r \ wf r" by (erule finite_acyclic_wf) (simp add: acyclic_irrefl) lemma finite_irreflp_transp_imp_wfp: "finite {(x, y). p x y} \ irreflp p \ transp p \ wfP p" using finite_irrefl_trans_imp_wf[of "{(x, y). p x y}"] unfolding wfP_def transp_def irreflp_def trans_def irrefl_def mem_Collect_eq prod.case by assumption lemma wf_infinite_down_chain_compatible: assumes wf_R: "wf R" and inf_chain_RS: "\i. (f (Suc i), f i) \ R \ S" and O_subset: "R O S \ R" shows "\k. \i. (f (Suc (i + k)), f (i + k)) \ S" proof (rule ccontr) assume "\k. \i. (f (Suc (i + k)), f (i + k)) \ S" hence "\k. \i. (f (Suc (i + k)), f (i + k)) \ S" by blast hence "\k. \i > k. (f (Suc i), f i) \ S" by (metis add.commute add_Suc less_add_Suc1) hence "\k. \i > k. (f (Suc i), f i) \ R" using inf_chain_RS by blast hence "\i > k. (f (Suc i), f i) \ R \ (\j > k. (f (Suc j), f j) \ R \ j \ i)" for k using wf_eq_minimal[THEN iffD1, OF wf_less, rule_format, of _ "{i. i > k \ (f (Suc i), f i) \ R}", simplified] by (meson not_less) then obtain j_of where j_of_gt: "\k. j_of k > k" and j_of_in_R: "\k. (f (Suc (j_of k)), f (j_of k)) \ R" and j_of_min: "\k. \j > k. (f (Suc j), f j) \ R \ j \ j_of k" by moura have j_of_min_s: "\k j. j > k \ j < j_of k \ (f (Suc j), f j) \ S" using j_of_min inf_chain_RS by fastforce define g :: "nat \ 'a" where "\k. g k = f (Suc ((j_of ^^ k) 0))" have between_g[simplified]: "(f ((j_of ^^ (Suc i)) 0), f (Suc ((j_of ^^ i) 0))) \ S\<^sup>*" for i proof (rule f_rev_chain_in_rtrancl; clarify?) show "Suc ((j_of ^^ i) 0) \ (j_of ^^ Suc i) 0" using j_of_gt by (simp add: Suc_leI) next fix ia assume ia: "ia \ {Suc ((j_of ^^ i) 0)..<(j_of ^^ Suc i) 0}" have ia_gt: "ia > (j_of ^^ i) 0" using ia by auto have ia_lt: "ia < j_of ((j_of ^^ i) 0)" using ia by auto show "(f (Suc ia), f ia) \ S" by (rule j_of_min_s[OF ia_gt ia_lt]) qed have "\i. (g (Suc i), g i) \ R" unfolding g_def funpow.simps comp_def by (rule subsetD[OF relcomp_subset_left_imp_relcomp_trancl_subset_left[OF O_subset]]) (rule relcompI[OF j_of_in_R between_g]) moreover have "\f. \i. (f (Suc i), f i) \ R" using wf_R[unfolded wf_iff_no_infinite_down_chain] by blast ultimately show False by blast qed subsection \Wellorders\ lemma (in wellorder) exists_minimal: fixes x :: 'a assumes "P x" shows "\x. P x \ (\y. P y \ y \ x)" using assms by (auto intro: LeastI Least_le) subsection \Lists\ lemma rev_induct2[consumes 1, case_names Nil snoc]: "length xs = length ys \ P [] [] \ (\x xs y ys. length xs = length ys \ P xs ys \ P (xs @ [x]) (ys @ [y])) \ P xs ys" proof (induct xs arbitrary: ys rule: rev_induct) case (snoc x xs ys) thus ?case by (induct ys rule: rev_induct) simp_all qed auto lemma hd_in_set: "length xs \ 0 \ hd xs \ set xs" by (cases xs) auto lemma in_lists_iff_set: "xs \ lists A \ set xs \ A" by fast lemma butlast_append_Cons[simp]: "butlast (xs @ y # ys) = xs @ butlast (y # ys)" using butlast_append[of xs "y # ys", simplified] by simp lemma rev_in_lists[simp]: "rev xs \ lists A \ xs \ lists A" by auto lemma hd_le_sum_list: fixes xs :: "'a::ordered_ab_semigroup_monoid_add_imp_le list" assumes "xs \ []" and "\i < length xs. xs ! i \ 0" shows "hd xs \ sum_list xs" using assms by (induct xs rule: rev_induct, simp_all, metis add_cancel_right_left add_increasing2 hd_append2 lessI less_SucI list.sel(1) nth_append nth_append_length order_refl self_append_conv2 sum_list.Nil) lemma sum_list_ge_length_times: fixes a :: "'a::{ordered_ab_semigroup_add,semiring_1}" assumes "\i < length xs. xs ! i \ a" shows "sum_list xs \ of_nat (length xs) * a" using assms proof (induct xs) case (Cons x xs) note ih = this(1) and xxs_i_ge_a = this(2) have xs_i_ge_a: "\i < length xs. xs ! i \ a" using xxs_i_ge_a by auto have "x \ a" using xxs_i_ge_a by auto thus ?case using ih[OF xs_i_ge_a] by (simp add: ring_distribs ordered_ab_semigroup_add_class.add_mono) qed auto lemma prod_list_nonneg: fixes xs :: "('a :: {ordered_semiring_0,linordered_nonzero_semiring}) list" assumes "\x. x \ set xs \ x \ 0" shows "prod_list xs \ 0" using assms by (induct xs) auto lemma zip_append_0_upt: "zip (xs @ ys) [0.. 0" and len_eq: "length xs = length ys" shows "zip xs ys = zip (butlast xs) (butlast ys) @ [(last xs, last ys)]" using len_eq len_gt0 by (induct rule: list_induct2) auto subsection \Extended Natural Numbers\ lemma the_enat_0[simp]: "the_enat 0 = 0" by (simp add: zero_enat_def) lemma the_enat_1[simp]: "the_enat 1 = 1" by (simp add: one_enat_def) lemma enat_le_minus_1_imp_lt: "m \ n - 1 \ n \ \ \ n \ 0 \ m < n" for m n :: enat by (cases m; cases n; simp add: zero_enat_def one_enat_def) lemma enat_diff_diff_eq: "k - m - n = k - (m + n)" for k m n :: enat by (cases k; cases m; cases n) auto lemma enat_sub_add_same[intro]: "n \ m \ m = m - n + n" for m n :: enat by (cases m; cases n) auto lemma enat_the_enat_iden[simp]: "n \ \ \ enat (the_enat n) = n" by auto lemma the_enat_minus_nat: "m \ \ \ the_enat (m - enat n) = the_enat m - n" by auto lemma enat_the_enat_le: "enat (the_enat x) \ x" by (cases x; simp) lemma enat_the_enat_minus_le: "enat (the_enat (x - y)) \ x" by (cases x; cases y; simp) lemma enat_le_imp_minus_le: "k \ m \ k - n \ m" for k m n :: enat by (metis Groups.add_ac(2) enat_diff_diff_eq enat_ord_simps(3) enat_sub_add_same enat_the_enat_iden enat_the_enat_minus_le idiff_0_right idiff_infinity idiff_infinity_right order_trans_rules(23) plus_enat_simps(3)) lemma add_diff_assoc2_enat: "m \ n \ m - n + p = m + p - n" for m n p :: enat by (cases m; cases n; cases p; auto) lemma enat_mult_minus_distrib: "enat x * (y - z) = enat x * y - enat x * z" by (cases y; cases z; auto simp: enat_0 right_diff_distrib') subsection \Multisets\ lemma add_mset_lt_left_lt: "a < b \ add_mset a A < add_mset b A" unfolding less_multiset\<^sub>H\<^sub>O by auto lemma add_mset_le_left_le: "a \ b \ add_mset a A \ add_mset b A" for a :: "'a :: linorder" unfolding less_multiset\<^sub>H\<^sub>O by auto lemma add_mset_lt_right_lt: "A < B \ add_mset a A < add_mset a B" unfolding less_multiset\<^sub>H\<^sub>O by auto lemma add_mset_le_right_le: "A \ B \ add_mset a A \ add_mset a B" unfolding less_multiset\<^sub>H\<^sub>O by auto lemma add_mset_lt_lt_lt: assumes a_lt_b: "a < b" and A_le_B: "A < B" shows "add_mset a A < add_mset b B" by (rule less_trans[OF add_mset_lt_left_lt[OF a_lt_b] add_mset_lt_right_lt[OF A_le_B]]) lemma add_mset_lt_lt_le: "a < b \ A \ B \ add_mset a A < add_mset b B" using add_mset_lt_lt_lt le_neq_trans by fastforce lemma add_mset_lt_le_lt: "a \ b \ A < B \ add_mset a A < add_mset b B" for a :: "'a :: linorder" using add_mset_lt_lt_lt by (metis add_mset_lt_right_lt le_less) lemma add_mset_le_le_le: fixes a :: "'a :: linorder" assumes a_le_b: "a \ b" and A_le_B: "A \ B" shows "add_mset a A \ add_mset b B" by (rule order.trans[OF add_mset_le_left_le[OF a_le_b] add_mset_le_right_le[OF A_le_B]]) declare filter_eq_replicate_mset [simp] image_mset_subseteq_mono [intro] lemma nonempty_subseteq_mset_eq_singleton: "M \ {#} \ M \# {#x#} \ M = {#x#}" by (cases M) (auto dest: subset_mset.diff_add) lemma nonempty_subseteq_mset_iff_singleton: "(M \ {#} \ M \# {#x#} \ P) \ M = {#x#} \ P" by (cases M) (auto dest: subset_mset.diff_add) lemma count_gt_imp_in_mset[intro]: "count M x > n \ x \# M" using count_greater_zero_iff by fastforce lemma size_lt_imp_ex_count_lt: "size M < size N \ \x \# N. count M x < count N x" by (metis count_eq_zero_iff leD not_le_imp_less not_less_zero size_mset_mono subseteq_mset_def) lemma filter_filter_mset[simp]: "{#x \# {#x \# M. Q x#}. P x#} = {#x \# M. P x \ Q x#}" by (induct M) auto lemma size_filter_unsat_elem: assumes "x \# M" and "\ P x" shows "size {#x \# M. P x#} < size M" proof - have "size (filter_mset P M) \ size M" using assms by (metis add.right_neutral add_diff_cancel_left' count_filter_mset mem_Collect_eq multiset_partition nonempty_has_size set_mset_def size_union) then show ?thesis by (meson leD nat_neq_iff size_filter_mset_lesseq) qed lemma size_filter_ne_elem: "x \# M \ size {#y \# M. y \ x#} < size M" by (simp add: size_filter_unsat_elem[of x M "\y. y \ x"]) lemma size_eq_ex_count_lt: assumes sz_m_eq_n: "size M = size N" and m_eq_n: "M \ N" shows "\x. count M x < count N x" proof - obtain x where "count M x \ count N x" using m_eq_n by (meson multiset_eqI) moreover { assume "count M x < count N x" hence ?thesis by blast } moreover { assume cnt_x: "count M x > count N x" have "size {#y \# M. y = x#} + size {#y \# M. y \ x#} = size {#y \# N. y = x#} + size {#y \# N. y \ x#}" using sz_m_eq_n multiset_partition by (metis size_union) hence sz_m_minus_x: "size {#y \# M. y \ x#} < size {#y \# N. y \ x#}" using cnt_x by simp then obtain y where "count {#y \# M. y \ x#} y < count {#y \# N. y \ x#} y" using size_lt_imp_ex_count_lt[OF sz_m_minus_x] by blast hence "count M y < count N y" by (metis count_filter_mset less_asym) hence ?thesis by blast } ultimately show ?thesis by fastforce qed lemma count_image_mset_lt_imp_lt_raw: assumes "finite A" and "A = set_mset M \ set_mset N" and "count (image_mset f M) b < count (image_mset f N) b" shows "\x. f x = b \ count M x < count N x" using assms proof (induct A arbitrary: M N b rule: finite_induct) case (insert x F) note fin = this(1) and x_ni_f = this(2) and ih = this(3) and x_f_eq_m_n = this(4) and cnt_b = this(5) let ?Ma = "{#y \# M. y \ x#}" let ?Mb = "{#y \# M. y = x#}" let ?Na = "{#y \# N. y \ x#}" let ?Nb = "{#y \# N. y = x#}" have m_part: "M = ?Mb + ?Ma" and n_part: "N = ?Nb + ?Na" using multiset_partition by blast+ have f_eq_ma_na: "F = set_mset ?Ma \ set_mset ?Na" using x_f_eq_m_n x_ni_f by auto show ?case proof (cases "count (image_mset f ?Ma) b < count (image_mset f ?Na) b") case cnt_ba: True obtain xa where "f xa = b" and "count ?Ma xa < count ?Na xa" using ih[OF f_eq_ma_na cnt_ba] by blast thus ?thesis by (metis count_filter_mset not_less0) next case cnt_ba: False have fx_eq_b: "f x = b" using cnt_b cnt_ba by (subst (asm) m_part, subst (asm) n_part, auto, presburger) moreover have "count M x < count N x" using cnt_b cnt_ba by (subst (asm) m_part, subst (asm) n_part, auto simp: fx_eq_b) ultimately show ?thesis by blast qed qed auto lemma count_image_mset_lt_imp_lt: assumes cnt_b: "count (image_mset f M) b < count (image_mset f N) b" shows "\x. f x = b \ count M x < count N x" by (rule count_image_mset_lt_imp_lt_raw[of "set_mset M \ set_mset N", OF _ refl cnt_b]) auto lemma count_image_mset_le_imp_lt_raw: assumes "finite A" and "A = set_mset M \ set_mset N" and "count (image_mset f M) (f a) + count N a < count (image_mset f N) (f a) + count M a" shows "\b. f b = f a \ count M b < count N b" using assms proof (induct A arbitrary: M N rule: finite_induct) case (insert x F) note fin = this(1) and x_ni_f = this(2) and ih = this(3) and x_f_eq_m_n = this(4) and cnt_lt = this(5) let ?Ma = "{#y \# M. y \ x#}" let ?Mb = "{#y \# M. y = x#}" let ?Na = "{#y \# N. y \ x#}" let ?Nb = "{#y \# N. y = x#}" have m_part: "M = ?Mb + ?Ma" and n_part: "N = ?Nb + ?Na" using multiset_partition by blast+ have f_eq_ma_na: "F = set_mset ?Ma \ set_mset ?Na" using x_f_eq_m_n x_ni_f by auto show ?case proof (cases "f x = f a") case fx_ne_fa: False have cnt_fma_fa: "count (image_mset f ?Ma) (f a) = count (image_mset f M) (f a)" using fx_ne_fa by (subst (2) m_part) auto have cnt_fna_fa: "count (image_mset f ?Na) (f a) = count (image_mset f N) (f a)" using fx_ne_fa by (subst (2) n_part) auto have cnt_ma_a: "count ?Ma a = count M a" using fx_ne_fa by (subst (2) m_part) auto have cnt_na_a: "count ?Na a = count N a" using fx_ne_fa by (subst (2) n_part) auto obtain b where fb_eq_fa: "f b = f a" and cnt_b: "count ?Ma b < count ?Na b" using ih[OF f_eq_ma_na] cnt_lt unfolding cnt_fma_fa cnt_fna_fa cnt_ma_a cnt_na_a by blast have fx_ne_fb: "f x \ f b" using fb_eq_fa fx_ne_fa by simp have cnt_ma_b: "count ?Ma b = count M b" using fx_ne_fb by (subst (2) m_part) auto have cnt_na_b: "count ?Na b = count N b" using fx_ne_fb by (subst (2) n_part) auto show ?thesis using fb_eq_fa cnt_b unfolding cnt_ma_b cnt_na_b by blast next case fx_eq_fa: True show ?thesis proof (cases "x = a") case x_eq_a: True have "count (image_mset f ?Ma) (f a) + count ?Na a < count (image_mset f ?Na) (f a) + count ?Ma a" using cnt_lt x_eq_a by (subst (asm) (1 2) m_part, subst (asm) (1 2) n_part, auto) thus ?thesis using ih[OF f_eq_ma_na] by (metis count_filter_mset nat_neq_iff) next case x_ne_a: False show ?thesis proof (cases "count M x < count N x") case True thus ?thesis using fx_eq_fa by blast next case False hence cnt_x: "count M x \ count N x" by fastforce have "count M x + count (image_mset f ?Ma) (f a) + count ?Na a < count N x + count (image_mset f ?Na) (f a) + count ?Ma a" using cnt_lt x_ne_a fx_eq_fa by (subst (asm) (1 2) m_part, subst (asm) (1 2) n_part, auto) hence "count (image_mset f ?Ma) (f a) + count ?Na a < count (image_mset f ?Na) (f a) + count ?Ma a" using cnt_x by linarith thus ?thesis using ih[OF f_eq_ma_na] by (metis count_filter_mset nat_neq_iff) qed qed qed qed auto lemma count_image_mset_le_imp_lt: assumes "count (image_mset f M) (f a) \ count (image_mset f N) (f a)" and "count M a > count N a" shows "\b. f b = f a \ count M b < count N b" using assms by (auto intro: count_image_mset_le_imp_lt_raw[of "set_mset M \ set_mset N"]) lemma Max_in_mset: "M \ {#} \ Max_mset M \# M" by simp lemma Max_lt_imp_lt_mset: assumes n_nemp: "N \ {#}" and max: "Max_mset M < Max_mset N" (is "?max_M < ?max_N") shows "M < N" proof (cases "M = {#}") case m_nemp: False have max_n_in_n: "?max_N \# N" using n_nemp by simp have max_n_nin_m: "?max_N \# M" using max Max_ge leD by auto have "M \ N" using max by auto moreover { fix y assume "count N y < count M y" hence "y \# M" by blast hence "?max_M \ y" by simp hence "?max_N > y" using max by auto hence "\x > y. count M x < count N x" using max_n_nin_m max_n_in_n by fastforce } ultimately show ?thesis unfolding less_multiset\<^sub>H\<^sub>O by blast qed (auto simp: n_nemp) lemma fold_mset_singleton[simp]: "fold_mset f z {#x#} = f x z" by (simp add: fold_mset_def) end diff --git a/thys/Perron_Frobenius/Perron_Frobenius_Irreducible.thy b/thys/Perron_Frobenius/Perron_Frobenius_Irreducible.thy --- a/thys/Perron_Frobenius/Perron_Frobenius_Irreducible.thy +++ b/thys/Perron_Frobenius/Perron_Frobenius_Irreducible.thy @@ -1,1378 +1,1378 @@ (* Author: Thiemann *) subsection \The Perron Frobenius Theorem for Irreducible Matrices\ theory Perron_Frobenius_Irreducible imports Perron_Frobenius Roots_Unity Rank_Nullity_Theorem.Miscellaneous (* for scalar-matrix-multiplication, this import is incompatible with field_simps, ac_simps *) begin lifting_forget vec.lifting lifting_forget mat.lifting lifting_forget poly.lifting lemma charpoly_of_real: "charpoly (map_matrix complex_of_real A) = map_poly of_real (charpoly A)" by (transfer_hma rule: of_real_hom.char_poly_hom) context includes lifting_syntax begin lemma HMA_M_smult[transfer_rule]: "((=) ===> HMA_M ===> HMA_M) (\\<^sub>m) ((*k))" unfolding smult_mat_def unfolding rel_fun_def HMA_M_def from_hma\<^sub>m_def by (auto simp: matrix_scalar_mult_def) end lemma order_charpoly_smult: fixes A :: "complex ^ 'n ^ 'n" assumes k: "k \ 0" shows "order x (charpoly (k *k A)) = order (x / k) (charpoly A)" by (transfer fixing: k, rule order_char_poly_smult[OF _ k]) (* use field, since the *k-lemmas have been stated for fields *) lemma smult_eigen_vector: fixes a :: "'a :: field" assumes "eigen_vector A v x" shows "eigen_vector (a *k A) v (a * x)" proof - from assms[unfolded eigen_vector_def] have v: "v \ 0" and id: "A *v v = x *s v" by auto from arg_cong[OF id, of "(*s) a"] have id: "(a *k A) *v v = (a * x) *s v" unfolding scalar_matrix_vector_assoc by simp thus "eigen_vector (a *k A) v (a * x)" using v unfolding eigen_vector_def by auto qed lemma smult_eigen_value: fixes a :: "'a :: field" assumes "eigen_value A x" shows "eigen_value (a *k A) (a * x)" using assms smult_eigen_vector[of A _ x a] unfolding eigen_value_def by blast locale fixed_mat = fixes A :: "'a :: zero ^ 'n ^ 'n" begin definition G :: "'n rel" where "G = { (i,j). A $ i $ j \ 0}" definition irreducible :: bool where "irreducible = (UNIV \ G^+)" end lemma G_transpose: "fixed_mat.G (transpose A) = ((fixed_mat.G A))^-1" unfolding fixed_mat.G_def by (force simp: transpose_def) lemma G_transpose_trancl: "(fixed_mat.G (transpose A))^+ = ((fixed_mat.G A)^+)^-1" unfolding G_transpose trancl_converse by auto locale pf_nonneg_mat = fixed_mat A for A :: "'a :: linordered_idom ^ 'n ^ 'n" + assumes non_neg_mat: "non_neg_mat A" begin lemma nonneg: "A $ i $ j \ 0" using non_neg_mat unfolding non_neg_mat_def elements_mat_h_def by auto lemma nonneg_matpow: "matpow A n $ i $ j \ 0" by (induct n arbitrary: i j, insert nonneg, auto intro!: sum_nonneg simp: matrix_matrix_mult_def mat_def) lemma G_relpow_matpow_pos: "(i,j) \ G ^^ n \ matpow A n $ i $ j > 0" proof (induct n arbitrary: i j) case (0 i) thus ?case by (auto simp: mat_def) next case (Suc n i j) from Suc(2) have "(i,j) \ G ^^ n O G" by (simp add: relpow_commute) then obtain k where ik: "A $ k $ j \ 0" and kj: "(i, k) \ G ^^ n" by (auto simp: G_def) from ik nonneg[of k j] have ik: "A $ k $ j > 0" by auto from Suc(1)[OF kj] have IH: "matpow A n $h i $h k > 0" . thus ?case using ik by (auto simp: nonneg_matpow nonneg matrix_matrix_mult_def intro!: sum_pos2[of _ k] mult_nonneg_nonneg) qed lemma matpow_mono: assumes B: "\ i j. B $ i $ j \ A $ i $ j" shows "matpow B n $ i $ j \ matpow A n $ i $ j" proof (induct n arbitrary: i j) case (Suc n i j) thus ?case using B nonneg_matpow[of n] nonneg by (auto simp: matrix_matrix_mult_def intro!: sum_mono mult_mono') qed simp lemma matpow_sum_one_mono: "matpow (A + mat 1) (n + k) $ i $ j \ matpow (A + mat 1) n $ i $ j" proof (induct k) case (Suc k) have "(matpow (A + mat 1) (n + k) ** A) $h i $h j \ 0" unfolding matrix_matrix_mult_def using order.trans[OF nonneg_matpow matpow_mono[of "A + mat 1" "n + k"]] by (auto intro!: sum_nonneg mult_nonneg_nonneg nonneg simp: mat_def) thus ?case using Suc by (simp add: matrix_add_ldistrib matrix_mul_rid) qed simp lemma G_relpow_matpow_pos_ge: assumes "(i,j) \ G ^^ m" "n \ m" shows "matpow (A + mat 1) n $ i $ j > 0" proof - from assms(2) obtain k where n: "n = m + k" using le_Suc_ex by blast have "0 < matpow A m $ i $ j" by (rule G_relpow_matpow_pos[OF assms(1)]) also have "\ \ matpow (A + mat 1) m $ i $ j" by (rule matpow_mono, auto simp: mat_def) also have "\ \ matpow (A + mat 1) n $ i $ j" unfolding n using matpow_sum_one_mono . finally show ?thesis . qed end locale perron_frobenius = pf_nonneg_mat A for A :: "real ^ 'n ^ 'n" + assumes irr: irreducible begin definition N where "N = (SOME N. \ ij. \ n \ N. ij \ G ^^ n)" lemma N: "\ n \ N. ij \ G ^^ n" proof - { fix ij have "ij \ G^+" using irr[unfolded irreducible_def] by auto from this[unfolded trancl_power] have "\ n. ij \ G ^^ n" by blast } hence "\ ij. \ n. ij \ G ^^ n" by auto from choice[OF this] obtain f where f: "\ ij. ij \ G ^^ (f ij)" by auto define N where N: "N = Max (range f)" { fix ij from f[of ij] have "ij \ G ^^ f ij" . moreover have "f ij \ N" unfolding N by (rule Max_ge, auto) ultimately have "\ n \ N. ij \ G ^^ n" by blast } note main = this let ?P = "\ N. \ ij. \ n \ N. ij \ G ^^ n" from main have "?P N" by blast from someI[of ?P, OF this, folded N_def] show ?thesis by blast qed lemma irreducible_matpow_pos: assumes irreducible shows "matpow (A + mat 1) N $ i $ j > 0" proof - from N obtain n where n: "n \ N" and reach: "(i,j) \ G ^^ n" by auto show ?thesis by (rule G_relpow_matpow_pos_ge[OF reach n]) qed lemma pf_transpose: "perron_frobenius (transpose A)" proof show "fixed_mat.irreducible (transpose A)" unfolding fixed_mat.irreducible_def G_transpose_trancl using irr[unfolded irreducible_def] by auto qed (insert nonneg, auto simp: transpose_def non_neg_mat_def elements_mat_h_def) abbreviation le_vec :: "real ^ 'n \ real ^ 'n \ bool" where "le_vec x y \ (\ i. x $ i \ y $ i)" abbreviation lt_vec :: "real ^ 'n \ real ^ 'n \ bool" where "lt_vec x y \ (\ i. x $ i < y $ i)" definition "A1n = matpow (A + mat 1) N" lemmas A1n_pos = irreducible_matpow_pos[OF irr, folded A1n_def] definition r :: "real ^ 'n \ real" where "r x = Min { (A *v x) $ j / x $ j | j. x $ j \ 0 }" definition X :: "(real ^ 'n )set" where "X = { x . le_vec 0 x \ x \ 0 }" lemma nonneg_Ax: "x \ X \ le_vec 0 (A *v x)" unfolding X_def using nonneg by (auto simp: matrix_vector_mult_def intro!: sum_nonneg) lemma A_nonzero_fixed_i: "\ j. A $ i $ j \ 0" proof - from irr[unfolded irreducible_def] have "(i,i) \ G^+" by auto then obtain j where "(i,j) \ G" by (metis converse_tranclE) hence Aij: "A $ i $ j \ 0" unfolding G_def by auto thus ?thesis .. qed lemma A_nonzero_fixed_j: "\ i. A $ i $ j \ 0" proof - from irr[unfolded irreducible_def] have "(j,j) \ G^+" by auto then obtain i where "(i,j) \ G" by (cases, auto) hence Aij: "A $ i $ j \ 0" unfolding G_def by auto thus ?thesis .. qed lemma Ax_pos: assumes x: "lt_vec 0 x" shows "lt_vec 0 (A *v x)" proof fix i from A_nonzero_fixed_i[of i] obtain j where "A $ i $ j \ 0" by auto with nonneg[of i j] have A: "A $ i $ j > 0" by simp from x have "x $ j \ 0" for j by (auto simp: order.strict_iff_order) note nonneg = mult_nonneg_nonneg[OF nonneg[of i] this] have "(A *v x) $ i = (\j\UNIV. A $ i $ j * x $ j)" unfolding matrix_vector_mult_def by simp also have "\ = A $ i $ j * x $ j + (\j\UNIV - {j}. A $ i $ j * x $ j)" by (subst sum.remove, auto) also have "\ > 0 + 0" by (rule add_less_le_mono, insert A x[rule_format] nonneg, auto intro!: sum_nonneg mult_pos_pos) finally show "0 $ i < (A *v x) $ i" by simp qed lemma nonzero_Ax: assumes x: "x \ X" shows "A *v x \ 0" proof assume 0: "A *v x = 0" from x[unfolded X_def] have x: "le_vec 0 x" "x \ 0" by auto from x(2) obtain j where xj: "x $ j \ 0" by (metis vec_eq_iff zero_index) from A_nonzero_fixed_j[of j] obtain i where Aij: "A $ i $ j \ 0" by auto from arg_cong[OF 0, of "\ v. v $ i", unfolded matrix_vector_mult_def] have "0 = (\ k \ UNIV. A $h i $h k * x $h k)" by auto also have "\ = A $h i $h j * x $h j + (\ k \ UNIV - {j}. A $h i $h k * x $h k)" by (subst sum.remove[of _ j], auto) also have "\ > 0 + 0" by (rule add_less_le_mono, insert nonneg[of i] Aij x(1) xj, auto intro!: sum_nonneg mult_pos_pos simp: dual_order.not_eq_order_implies_strict) finally show False by simp qed lemma r_witness: assumes x: "x \ X" shows "\ j. x $ j > 0 \ r x = (A *v x) $ j / x $ j" proof - from x[unfolded X_def] have x: "le_vec 0 x" "x \ 0" by auto let ?A = "{ (A *v x) $ j / x $ j | j. x $ j \ 0 }" from x(2) obtain j where "x $ j \ 0" by (metis vec_eq_iff zero_index) hence empty: "?A \ {}" by auto from Min_in[OF _ this, folded r_def] obtain j where "x $ j \ 0" and rx: "r x = (A *v x) $ j / x $ j" by auto with x have "x $ j > 0" by (auto simp: dual_order.not_eq_order_implies_strict) with rx show ?thesis by auto qed lemma rx_nonneg: assumes x: "x \ X" shows "r x \ 0" proof - from x[unfolded X_def] have x: "le_vec 0 x" "x \ 0" by auto let ?A = "{ (A *v x) $ j / x $ j | j. x $ j \ 0 }" from r_witness[OF \x \ X\] have empty: "?A \ {}" by force show ?thesis unfolding r_def X_def proof (subst Min_ge_iff, force, use empty in force, intro ballI) fix y assume "y \ ?A" then obtain j where "y = (A *v x) $ j / x $ j" and "x $ j \ 0" by auto from nonneg_Ax[OF \x \ X\] this x show "0 \ y" by simp qed qed lemma rx_pos: assumes x: "lt_vec 0 x" shows "r x > 0" proof - from Ax_pos[OF x] have lt: "lt_vec 0 (A *v x)" . from x have x': "x \ X" unfolding X_def order.strict_iff_order by auto let ?A = "{ (A *v x) $ j / x $ j | j. x $ j \ 0 }" from r_witness[OF \x \ X\] have empty: "?A \ {}" by force show ?thesis unfolding r_def X_def proof (subst Min_gr_iff, force, use empty in force, intro ballI) fix y assume "y \ ?A" then obtain j where "y = (A *v x) $ j / x $ j" and "x $ j \ 0" by auto from lt this x show "0 < y" by simp qed qed lemma rx_le_Ax: assumes x: "x \ X" shows "le_vec (r x *s x) (A *v x)" proof (intro allI) fix i show "(r x *s x) $h i \ (A *v x) $h i" proof (cases "x $ i = 0") case True with nonneg_Ax[OF x] show ?thesis by auto next case False with x[unfolded X_def] have pos: "x $ i > 0" by (auto simp: dual_order.not_eq_order_implies_strict) from False have "(A *v x) $h i / x $ i \ { (A *v x) $ j / x $ j | j. x $ j \ 0 }" by auto hence "(A *v x) $h i / x $ i \ r x" unfolding r_def by simp hence "x $ i * r x \ x $ i * ((A *v x) $h i / x $ i)" unfolding mult_le_cancel_left_pos[OF pos] . also have "\ = (A *v x) $h i" using pos by simp finally show ?thesis by (simp add: ac_simps) qed qed lemma rho_le_x_Ax_imp_rho_le_rx: assumes x: "x \ X" and \: "le_vec (\ *s x) (A *v x)" shows "\ \ r x" proof - from r_witness[OF x] obtain j where rx: "r x = (A *v x) $ j / x $ j" and xj: "x $ j > 0" "x $ j \ 0" by auto from divide_right_mono[OF \[rule_format, of j] xj(2)] show ?thesis unfolding rx using xj by simp qed lemma rx_Max: assumes x: "x \ X" shows "r x = Sup { \ . le_vec (\ *s x) (A *v x) }" (is "_ = Sup ?S") proof - have "r x \ ?S" using rx_le_Ax[OF x] by auto moreover { fix y assume "y \ ?S" hence y: "le_vec (y *s x) (A *v x)" by auto from rho_le_x_Ax_imp_rho_le_rx[OF x this] have "y \ r x" . } ultimately show ?thesis by (metis (mono_tags, lifting) cSup_eq_maximum) qed lemma r_smult: assumes x: "x \ X" and a: "a > 0" shows "r (a *s x) = r x" unfolding r_def by (rule arg_cong[of _ _ Min], unfold vector_smult_distrib, insert a, simp) definition "X1 = (X \ {x. norm x = 1})" lemma bounded_X1: "bounded X1" unfolding bounded_iff X1_def by auto lemma closed_X1: "closed X1" proof - have X1: "X1 = {x. le_vec 0 x \ norm x = 1}" unfolding X1_def X_def by auto show ?thesis unfolding X1 by (intro closed_Collect_conj closed_Collect_all closed_Collect_le closed_Collect_eq, auto intro: continuous_intros) qed lemma compact_X1: "compact X1" using bounded_X1 closed_X1 by (simp add: compact_eq_bounded_closed) definition "pow_A_1 x = A1n *v x" lemma continuous_pow_A_1: "continuous_on R pow_A_1" unfolding pow_A_1_def continuous_on by (auto intro: tendsto_intros) definition "Y = pow_A_1 ` X1" lemma compact_Y: "compact Y" unfolding Y_def using compact_X1 continuous_pow_A_1[of X1] by (metis compact_continuous_image) lemma Y_pos_main: assumes y: "y \ pow_A_1 ` X" shows "y $ i > 0" proof - from y obtain x where x: "x \ X" and y: "y = pow_A_1 x" unfolding Y_def X1_def by auto from r_witness[OF x] obtain j where xj: "x $ j > 0" by auto from x[unfolded X_def] have xi: "x $ i \ 0" for i by auto have nonneg: "0 \ A1n $ i $ k * x $ k" for k using A1n_pos[of i k] xi[of k] by auto have "y $ i = (\j\UNIV. A1n $ i $ j * x $ j)" unfolding y pow_A_1_def matrix_vector_mult_def by simp also have "\ = A1n $ i $ j * x $ j + (\j\UNIV - {j}. A1n $ i $ j * x $ j)" by (subst sum.remove, auto) also have "\ > 0 + 0" by (rule add_less_le_mono, insert xj A1n_pos nonneg, auto intro!: sum_nonneg mult_pos_pos simp: dual_order.not_eq_order_implies_strict) finally show ?thesis by simp qed lemma Y_pos: assumes y: "y \ Y" shows "y $ i > 0" using Y_pos_main[of y i] y unfolding Y_def X1_def by auto lemma Y_nonzero: assumes y: "y \ Y" shows "y $ i \ 0" using Y_pos[OF y, of i] by auto definition r' :: "real ^ 'n \ real" where "r' x = Min (range (\ j. (A *v x) $ j / x $ j))" lemma r'_r: assumes x: "x \ Y" shows "r' x = r x" unfolding r'_def r_def proof (rule arg_cong[of _ _ Min]) have "range (\j. (A *v x) $ j / x $ j) \ {(A *v x) $ j / x $ j |j. x $ j \ 0}" (is "?L \ ?R") proof fix y assume "y \ ?L" then obtain j where "y = (A *v x) $ j / x $ j" by auto with Y_pos[OF x, of j] show "y \ ?R" by auto qed moreover have "?L \ ?R" by auto ultimately show "?L = ?R" by blast qed lemma continuous_Y_r: "continuous_on Y r" proof - have *: "(\ y \ Y. P y (r y)) = (\ y \ Y. P y (r' y))" for P using r'_r by auto have "continuous_on Y r = continuous_on Y r'" by (rule continuous_on_cong[OF refl r'_r[symmetric]]) also have \ unfolding continuous_on r'_def using Y_nonzero by (auto intro!: tendsto_Min tendsto_intros) finally show ?thesis . qed lemma X1_nonempty: "X1 \ {}" proof - define x where "x = ((\ i. if i = undefined then 1 else 0) :: real ^ 'n)" { assume "x = 0" from arg_cong[OF this, of "\ x. x $ undefined"] have False unfolding x_def by auto } hence x: "x \ 0" by auto moreover have "le_vec 0 x" unfolding x_def by auto moreover have "norm x = 1" unfolding norm_vec_def L2_set_def by (auto, subst sum.remove[of _ undefined], auto simp: x_def) ultimately show ?thesis unfolding X1_def X_def by auto qed lemma Y_nonempty: "Y \ {}" unfolding Y_def using X1_nonempty by auto definition z where "z = (SOME z. z \ Y \ (\ y \ Y. r y \ r z))" abbreviation "sr \ r z" lemma z: "z \ Y" and sr_max_Y: "\ y. y \ Y \ r y \ sr" proof - let ?P = "\ z. z \ Y \ (\ y \ Y. r y \ r z)" from continuous_attains_sup[OF compact_Y Y_nonempty continuous_Y_r] obtain y where "?P y" by blast from someI[of ?P, OF this, folded z_def] show "z \ Y" "\ y. y \ Y \ r y \ r z" by blast+ qed lemma Y_subset_X: "Y \ X" proof fix y assume "y \ Y" from Y_pos[OF this] show "y \ X" unfolding X_def by (auto simp: order.strict_iff_order) qed lemma zX: "z \ X" using z(1) Y_subset_X by auto lemma le_vec_mono_left: assumes B: "\ i j. B $ i $ j \ 0" and "le_vec x y" shows "le_vec (B *v x) (B *v y)" proof (intro allI) fix i show "(B *v x) $ i \ (B *v y) $ i" unfolding matrix_vector_mult_def using B[of i] assms(2) by (auto intro!: sum_mono mult_left_mono) qed lemma matpow_1_commute: "matpow (A + mat 1) n ** A = A ** matpow (A + mat 1) n" by (induct n, auto simp: matrix_add_rdistrib matrix_add_ldistrib matrix_mul_rid matrix_mul_lid matrix_mul_assoc[symmetric]) lemma A1n_commute: "A1n ** A = A ** A1n" unfolding A1n_def by (rule matpow_1_commute) lemma le_vec_pow_A_1: assumes le: "le_vec (rho *s x) (A *v x)" shows "le_vec (rho *s pow_A_1 x) (A *v pow_A_1 x)" proof - have "A1n $ i $ j \ 0" for i j using A1n_pos[of i j] by auto from le_vec_mono_left[OF this le] have "le_vec (A1n *v (rho *s x)) (A1n *v (A *v x))" . also have "A1n *v (A *v x) = (A1n ** A) *v x" by (simp add: matrix_vector_mul_assoc) also have "\ = A *v (A1n *v x)" unfolding A1n_commute by (simp add: matrix_vector_mul_assoc) also have "\ = A *v (pow_A_1 x)" unfolding pow_A_1_def .. also have "A1n *v (rho *s x) = rho *s (A1n *v x)" unfolding vector_smult_distrib .. also have "\ = rho *s pow_A_1 x" unfolding pow_A_1_def .. finally show "le_vec (rho *s pow_A_1 x) (A *v pow_A_1 x)" . qed lemma r_pow_A_1: assumes x: "x \ X" shows "r x \ r (pow_A_1 x)" proof - let ?y = "pow_A_1 x" have "?y \ pow_A_1 ` X" using x by auto from Y_pos_main[OF this] have y: "?y \ X" unfolding X_def by (auto simp: order.strict_iff_order) let ?A = "{\. le_vec (\ *s x) (A *v x)}" let ?B = "{\. le_vec (\ *s pow_A_1 x) (A *v pow_A_1 x)}" show ?thesis unfolding rx_Max[OF x] rx_Max[OF y] proof (rule cSup_mono) show "bdd_above ?B" using rho_le_x_Ax_imp_rho_le_rx[OF y] by fast show "?A \ {}" using rx_le_Ax[OF x] by auto fix rho assume "rho \ ?A" hence "le_vec (rho *s x) (A *v x)" by auto from le_vec_pow_A_1[OF this] have "rho \ ?B" by auto thus "\ rho' \ ?B. rho \ rho'" by auto qed qed lemma sr_max: assumes x: "x \ X" shows "r x \ sr" proof - let ?n = "norm x" define x' where "x' = inverse ?n *s x" from x[unfolded X_def] have x0: "x \ 0" by auto hence n: "?n > 0" by auto have x': "x' \ X1" "x' \ X" using x n unfolding X1_def X_def x'_def by (auto simp: norm_smult) have id: "r x = r x'" unfolding x'_def by (rule sym, rule r_smult[OF x], insert n, auto) define y where "y = pow_A_1 x'" from x' have y: "y \ Y" unfolding Y_def y_def by auto note id also have "r x' \ r y" using r_pow_A_1[OF x'(2)] unfolding y_def . also have "\ \ r z" using sr_max_Y[OF y] . finally show "r x \ r z" . qed lemma z_pos: "z $ i > 0" using Y_pos[OF z(1)] by auto lemma sr_pos: "sr > 0" by (rule rx_pos, insert z_pos, auto) context fixes u assumes u: "u \ X" and ru: "r u = sr" begin lemma sr_imp_eigen_vector_main: "sr *s u = A *v u" proof (rule ccontr) assume *: "sr *s u \ A *v u" let ?x = "A *v u - sr *s u" from * have 0: "?x \ 0" by auto let ?y = "pow_A_1 u" have "le_vec (sr *s u) (A *v u)" using rx_le_Ax[OF u] unfolding ru . hence le: "le_vec 0 ?x" by auto from 0 le have x: "?x \ X" unfolding X_def by auto have y_pos: "lt_vec 0 ?y" using Y_pos_main[of ?y] u by auto hence y: "?y \ X" unfolding X_def by (auto simp: order.strict_iff_order) from Y_pos_main[of "pow_A_1 ?x"] x have "lt_vec 0 (pow_A_1 ?x)" by auto hence lt: "lt_vec (sr *s ?y) (A *v ?y)" unfolding pow_A_1_def matrix_vector_right_distrib_diff matrix_vector_mul_assoc A1n_commute vector_smult_distrib by simp let ?f = "(\ i. (A *v ?y - sr *s ?y) $ i / ?y $ i)" let ?U = "UNIV :: 'n set" define eps where "eps = Min (?f ` ?U)" have U: "finite (?f ` ?U)" "?f ` ?U \ {}" by auto have eps: "eps > 0" unfolding eps_def Min_gr_iff[OF U] using lt sr_pos y_pos by auto have le: "le_vec ((sr + eps) *s ?y) (A *v ?y)" proof fix i have "((sr + eps) *s ?y) $ i = sr * ?y $ i + eps * ?y $ i" by (simp add: comm_semiring_class.distrib) also have "\ \ sr * ?y $ i + ?f i * ?y $ i" proof (rule add_left_mono[OF mult_right_mono]) show "0 \ ?y $ i" using y_pos[rule_format, of i] by auto show "eps \ ?f i" unfolding eps_def by (rule Min_le, auto) qed also have "\ = (A *v ?y) $ i" using sr_pos y_pos[rule_format, of i] by simp finally show "((sr + eps) *s ?y) $ i \ (A *v ?y) $ i" . qed from rho_le_x_Ax_imp_rho_le_rx[OF y le] have "r ?y \ sr + eps" . with sr_max[OF y] eps show False by auto qed lemma sr_imp_eigen_vector: "eigen_vector A u sr" unfolding eigen_vector_def sr_imp_eigen_vector_main using u unfolding X_def by auto lemma sr_u_pos: "lt_vec 0 u" proof - let ?y = "pow_A_1 u" define n where "n = N" define c where "c = (sr + 1)^N" have c: "c > 0" using sr_pos unfolding c_def by auto have "lt_vec 0 ?y" using Y_pos_main[of ?y] u by auto also have "?y = A1n *v u" unfolding pow_A_1_def .. also have "\ = c *s u" unfolding c_def A1n_def n_def[symmetric] proof (induct n) case (Suc n) then show ?case by (simp add: matrix_vector_mul_assoc[symmetric] algebra_simps vec.scale sr_imp_eigen_vector_main[symmetric]) qed auto finally have lt: "lt_vec 0 (c *s u)" . have "0 < u $ i" for i using lt[rule_format, of i] c by simp (metis zero_less_mult_pos) thus "lt_vec 0 u" by simp qed end lemma eigen_vector_z_sr: "eigen_vector A z sr" using sr_imp_eigen_vector[OF zX refl] by auto lemma eigen_value_sr: "eigen_value A sr" using eigen_vector_z_sr unfolding eigen_value_def by auto abbreviation "c \ complex_of_real" abbreviation "cA \ map_matrix c A" abbreviation "norm_v \ map_vector (norm :: complex \ real)" lemma norm_v_ge_0: "le_vec 0 (norm_v v)" by (auto simp: map_vector_def) lemma norm_v_eq_0: "norm_v v = 0 \ v = 0" by (auto simp: map_vector_def vec_eq_iff) lemma cA_index: "cA $ i $ j = c (A $ i $ j)" unfolding map_matrix_def map_vector_def by simp lemma norm_cA[simp]: "norm (cA $ i $ j) = A $ i $ j" using nonneg[of i j] by (simp add: cA_index) context fixes \ v assumes ev: "eigen_vector cA v \" begin lemma evD: "\ *s v = cA *v v" "v \ 0" using ev[unfolded eigen_vector_def] by auto lemma ev_alpha_norm_v: "norm_v (\ *s v) = (norm \ *s norm_v v)" by (auto simp: map_vector_def norm_mult vec_eq_iff) lemma ev_A_norm_v: "norm_v (cA *v v) $ j \ (A *v norm_v v) $ j" proof - have "norm_v (cA *v v) $ j = norm (\i\UNIV. cA $ j $ i * v $ i)" unfolding map_vector_def by (simp add: matrix_vector_mult_def) also have "\ \ (\i\UNIV. norm (cA $ j $ i * v $ i))" by (rule norm_sum) also have "\ = (\i\UNIV. A $ j $ i * norm_v v $ i)" by (rule sum.cong[OF refl], auto simp: norm_mult map_vector_def) also have "\ = (A *v norm_v v) $ j" by (simp add: matrix_vector_mult_def) finally show ?thesis . qed lemma ev_le_vec: "le_vec (norm \ *s norm_v v) (A *v norm_v v)" using arg_cong[OF evD(1), of norm_v, unfolded ev_alpha_norm_v] ev_A_norm_v by auto lemma norm_v_X: "norm_v v \ X" using norm_v_ge_0[of v] evD(2) norm_v_eq_0[of v] unfolding X_def by auto lemma ev_inequalities: "norm \ \ r (norm_v v)" "r (norm_v v) \ sr" proof - have v: "norm_v v \ X" by (rule norm_v_X) from rho_le_x_Ax_imp_rho_le_rx[OF v ev_le_vec] show "norm \ \ r (norm_v v)" . from sr_max[OF v] show "r (norm_v v) \ sr" . qed lemma eigen_vector_norm_sr: "norm \ \ sr" using ev_inequalities by auto end lemma eigen_value_norm_sr: assumes "eigen_value cA \" shows "norm \ \ sr" using eigen_vector_norm_sr[of _ \] assms unfolding eigen_value_def by auto lemma le_vec_trans: "le_vec x y \ le_vec y u \ le_vec x u" using order.trans[of "x $ i" "y $ i" "u $ i" for i] by auto lemma eigen_vector_z_sr_c: "eigen_vector cA (map_vector c z) (c sr)" unfolding of_real_hom.eigen_vector_hom by (rule eigen_vector_z_sr) lemma eigen_value_sr_c: "eigen_value cA (c sr)" using eigen_vector_z_sr_c unfolding eigen_value_def by auto definition "w = perron_frobenius.z (transpose A)" lemma w: "transpose A *v w = sr *s w" "lt_vec 0 w" "perron_frobenius.sr (transpose A) = sr" proof - interpret t: perron_frobenius "transpose A" by (rule pf_transpose) from eigen_vector_z_sr_c t.eigen_vector_z_sr_c have ev: "eigen_value cA (c sr)" "eigen_value t.cA (c t.sr)" unfolding eigen_value_def by auto { fix x have "eigen_value (t.cA) x = eigen_value (transpose cA) x" unfolding map_matrix_def map_vector_def transpose_def by (auto simp: vec_eq_iff) also have "\ = eigen_value cA x" by (rule eigen_value_transpose) finally have "eigen_value (t.cA) x = eigen_value cA x" . } note ev_id = this with ev have ev: "eigen_value t.cA (c sr)" "eigen_value cA (c t.sr)" by auto from eigen_value_norm_sr[OF ev(2)] t.eigen_value_norm_sr[OF ev(1)] show id: "t.sr = sr" by auto from t.eigen_vector_z_sr[unfolded id, folded w_def] show "transpose A *v w = sr *s w" unfolding eigen_vector_def by auto from t.z_pos[folded w_def] show "lt_vec 0 w" by auto qed lemma c_cmod_id: "a \ \ \ Re a \ 0 \ c (cmod a) = a" by (auto simp: Reals_def) lemma pos_rowvector_mult_0: assumes lt: "lt_vec 0 x" and 0: "(rowvector x :: real ^ 'n ^ 'n) *v y = 0" (is "?x *v _ = 0") and le: "le_vec 0 y" shows "y = 0" proof - { fix i assume "y $ i \ 0" with le have yi: "y $ i > 0" by (auto simp: order.strict_iff_order) have "0 = (?x *v y) $ i" unfolding 0 by simp also have "\ = (\j\UNIV. x $ j * y $ j)" unfolding rowvector_def matrix_vector_mult_def by simp also have "\ > 0" by (rule sum_pos2[of _ i], insert yi lt le, auto intro!: mult_nonneg_nonneg simp: order.strict_iff_order) finally have False by simp } thus ?thesis by (auto simp: vec_eq_iff) qed lemma pos_matrix_mult_0: assumes le: "\ i j. B $ i $ j \ 0" and lt: "lt_vec 0 x" and 0: "B *v x = 0" shows "B = 0" proof - { fix i j assume "B $ i $ j \ 0" with le have gt: "B $ i $ j > 0" by (auto simp: order.strict_iff_order) have "0 = (B *v x) $ i" unfolding 0 by simp also have "\ = (\j\UNIV. B $ i $ j * x $ j)" unfolding matrix_vector_mult_def by simp also have "\ > 0" by (rule sum_pos2[of _ j], insert gt lt le, auto intro!: mult_nonneg_nonneg simp: order.strict_iff_order) finally have False by simp } thus "B = 0" unfolding vec_eq_iff by auto qed lemma eigen_value_smaller_matrix: assumes B: "\ i j. 0 \ B $ i $ j \ B $ i $ j \ A $ i $ j" and AB: "A \ B" and ev: "eigen_value (map_matrix c B) sigma" shows "cmod sigma < sr" proof - let ?B = "map_matrix c B" let ?sr = "spectral_radius ?B" define \ where "\ = ?sr" have "real_non_neg_mat ?B" unfolding real_non_neg_mat_def elements_mat_h_def by (auto simp: map_matrix_def map_vector_def B) from perron_frobenius[OF this, folded \_def] obtain x where ev_sr: "eigen_vector ?B x (c \)" and rnn: "real_non_neg_vec x" by auto define y where "y = norm_v x" from rnn have xy: "x = map_vector c y" unfolding real_non_neg_vec_def vec_elements_h_def y_def by (auto simp: map_vector_def vec_eq_iff c_cmod_id) from spectral_radius_max[OF ev, folded \_def] have sigma_sigma: "cmod sigma \ \" . from ev_sr[unfolded xy of_real_hom.eigen_vector_hom] have ev_B: "eigen_vector B y \" . from ev_B[unfolded eigen_vector_def] have ev_B': "B *v y = \ *s y" by auto have ypos: "y $ i \ 0" for i unfolding y_def by (auto simp: map_vector_def) from ev_B this have y: "y \ X" unfolding eigen_vector_def X_def by auto have BA: "(B *v y) $ i \ (A *v y) $ i" for i unfolding matrix_vector_mult_def vec_lambda_beta by (rule sum_mono, rule mult_right_mono, insert B ypos, auto) hence le_vec: "le_vec (\ *s y) (A *v y)" unfolding ev_B' by auto from rho_le_x_Ax_imp_rho_le_rx[OF y le_vec] have "\ \ r y" by auto also have "\ \ sr" using y by (rule sr_max) finally have sig_le_sr: "\ \ sr" . { assume "\ = sr" hence r_sr: "r y = sr" and sr_sig: "sr = \" using \\ \ r y\ \r y \ sr\ by auto from sr_u_pos[OF y r_sr] have pos: "lt_vec 0 y" . from sr_imp_eigen_vector[OF y r_sr] have ev': "eigen_vector A y sr" . have "(A - B) *v y = A *v y - B *v y" unfolding matrix_vector_mult_def by (auto simp: vec_eq_iff field_simps sum_subtractf) also have "A *v y = sr *s y" using ev'[unfolded eigen_vector_def] by auto also have "B *v y = sr *s y" unfolding ev_B' sr_sig .. finally have id: "(A - B) *v y = 0" by simp from pos_matrix_mult_0[OF _ pos id] assms(1-2) have False by auto } with sig_le_sr sigma_sigma show ?thesis by argo qed lemma charpoly_erase_mat_sr: "0 < poly (charpoly (erase_mat A i i)) sr" proof - let ?A = "erase_mat A i i" let ?pos = "poly (charpoly ?A) sr" { from A_nonzero_fixed_j[of i] obtain k where "A $ k $ i \ 0" by auto assume "A = ?A" hence "A $ k $ i = ?A $ k $ i" by simp also have "?A $ k $ i = 0" by (auto simp: erase_mat_def) also have "A $ k $ i \ 0" by fact finally have False by simp } hence AA: "A \ ?A" by auto have le: "0 \ ?A $ i $ j \ ?A $ i $ j \ A $ i $ j" for i j by (auto simp: erase_mat_def nonneg) note ev_small = eigen_value_smaller_matrix[OF le AA] { fix rho :: real assume "eigen_value ?A rho" hence ev: "eigen_value (map_matrix c ?A) (c rho)" unfolding eigen_value_def using of_real_hom.eigen_vector_hom[of ?A _ rho] by auto from ev_small[OF this] have "abs rho < sr" by auto } note ev_small_real = this have pos0: "?pos \ 0" using ev_small_real[of sr] by (auto simp: eigen_value_root_charpoly) { define p where "p = charpoly ?A" assume pos: "?pos < 0" hence neg: "poly p sr < 0" unfolding p_def by auto from degree_monic_charpoly[of ?A] have mon: "monic p" and deg: "degree p \ 0" unfolding p_def by auto let ?f = "poly p" have cont: "continuous_on {a..b} ?f" for a b by (auto intro: continuous_intros) from pos have le: "?f sr \ 0" by (auto simp: p_def) from mon have lc: "lead_coeff p > 0" by auto from poly_pinfty_ge[OF this deg, of 0] obtain z where lez: "\ x. z \ x \ 0 \ ?f x" by auto define y where "y = max z sr" have yr: "y \ sr" and "y \ z" unfolding y_def by auto from lez[OF this(2)] have y0: "?f y \ 0" . from IVT'[of ?f, OF le y0 yr cont] obtain x where ge: "x \ sr" and rt: "?f x = 0" unfolding p_def by auto hence "eigen_value ?A x" unfolding p_def by (simp add: eigen_value_root_charpoly) from ev_small_real[OF this] ge have False by auto } with pos0 show ?thesis by argo qed lemma multiplicity_sr_1: "order sr (charpoly A) = 1" proof - { assume "poly (pderiv (charpoly A)) sr = 0" hence "0 = poly (monom 1 1 * pderiv (charpoly A)) sr" by simp also have "\ = sum (\ i. poly (charpoly (erase_mat A i i)) sr) UNIV" unfolding pderiv_char_poly_erase_mat poly_sum .. also have "\ > 0" by (rule sum_pos, (force simp: charpoly_erase_mat_sr)+) finally have False by simp } hence nZ: "poly (pderiv (charpoly A)) sr \ 0" and nZ': "pderiv (charpoly A) \ 0" by auto from eigen_vector_z_sr have "eigen_value A sr" unfolding eigen_value_def .. from this[unfolded eigen_value_root_charpoly] have "poly (charpoly A) sr = 0" . hence "order sr (charpoly A) \ 0" unfolding order_root using nZ' by auto from order_pderiv[OF nZ' this] order_0I[OF nZ] show ?thesis by simp qed lemma sr_spectral_radius: "sr = spectral_radius cA" proof - from eigen_vector_z_sr_c have "eigen_value cA (c sr)" unfolding eigen_value_def by auto from spectral_radius_max[OF this] have sr: "sr \ spectral_radius cA" by auto with spectral_radius_ev[of cA] eigen_vector_norm_sr show ?thesis by force qed lemma le_vec_A_mu: assumes y: "y \ X" and le: "le_vec (A *v y) (mu *s y)" shows "sr \ mu" "lt_vec 0 y" "mu = sr \ A *v y = mu *s y \ mu = sr \ A *v y = mu *s y" proof - let ?w = "rowvector w" let ?w' = "columnvector w" have "?w ** A = transpose (transpose (?w ** A))" unfolding transpose_transpose by simp also have "transpose (?w ** A) = transpose A ** transpose ?w" by (rule matrix_transpose_mul) also have "transpose ?w = columnvector w" by (rule transpose_rowvector) also have "transpose A ** \ = columnvector (transpose A *v w)" unfolding dot_rowvector_columnvector[symmetric] .. also have "transpose A *v w = sr *s w" unfolding w by simp also have "transpose (columnvector \) = rowvector (sr *s w)" unfolding transpose_def columnvector_def rowvector_def vector_scalar_mult_def by auto finally have 1: "?w ** A = rowvector (sr *s w)" . have "sr *s (?w *v y) = ?w ** A *v y" unfolding 1 by (auto simp: rowvector_def vector_scalar_mult_def matrix_vector_mult_def vec_eq_iff sum_distrib_left mult.assoc) also have "\ = ?w *v (A *v y)" by (simp add: matrix_vector_mul_assoc) finally have eq1: "sr *s (rowvector w *v y) = rowvector w *v (A *v y)" . have "le_vec (rowvector w *v (A *v y)) (?w *v (mu *s y))" by (rule le_vec_mono_left[OF _ le], insert w(2), auto simp: rowvector_def order.strict_iff_order) also have "?w *v (mu *s y) = mu *s (?w *v y)" by (simp add: algebra_simps vec.scale) finally have le1: "le_vec (rowvector w *v (A *v y)) (mu *s (?w *v y))" . from le1[unfolded eq1[symmetric]] have 2: "le_vec (sr *s (?w *v y)) (mu *s (?w *v y))" . { from y obtain i where yi: "y $ i > 0" and y: "\ j. y $ j \ 0" unfolding X_def by (auto simp: order.strict_iff_order vec_eq_iff) from w(2) have wi: "w $ i > 0" and w: "\ j. w $ j \ 0" by (auto simp: order.strict_iff_order) have "(?w *v y) $ i > 0" using yi y wi w by (auto simp: matrix_vector_mult_def rowvector_def intro!: sum_pos2[of _ i] mult_nonneg_nonneg) moreover from 2[rule_format, of i] have "sr * (?w *v y) $ i \ mu * (?w *v y) $ i" by simp ultimately have "sr \ mu" by simp } thus *: "sr \ mu" . define cc where "cc = (mu + 1)^ N" define n where "n = N" from * sr_pos have mu: "mu \ 0" "mu > 0" by auto hence cc: "cc > 0" unfolding cc_def by simp from y have "pow_A_1 y \ pow_A_1 ` X" by auto from Y_pos_main[OF this] have lt: "0 < (A1n *v y) $ i" for i by (simp add: pow_A_1_def) have le: "le_vec (A1n *v y) (cc *s y)" unfolding cc_def A1n_def n_def[symmetric] proof (induct n) case (Suc n) let ?An = "matpow (A + mat 1) n" let ?mu = "(mu + 1)" have id': "matpow (A + mat 1) (Suc n) *v y = A *v (?An *v y) + ?An *v y" (is "?a = ?b + ?c") by (simp add: matrix_add_ldistrib matrix_mul_rid matrix_add_vect_distrib matpow_1_commute matrix_vector_mul_assoc[symmetric]) have "le_vec ?b (?mu^n *s (A *v y))" using le_vec_mono_left[OF nonneg Suc] by (simp add: algebra_simps vec.scale) moreover have "le_vec (?mu^n *s (A *v y)) (?mu^n *s (mu *s y))" using le mu by auto moreover have id: "?mu^n *s (mu *s y) = (?mu^n * mu) *s y" by simp from le_vec_trans[OF calculation[unfolded id]] have le1: "le_vec ?b ((?mu^n * mu) *s y)" . from Suc have le2: "le_vec ?c ((mu + 1) ^ n *s y)" . have le: "le_vec ?a ((?mu^n * mu) *s y + ?mu^n *s y)" unfolding id' using add_mono[OF le1[rule_format] le2[rule_format]] by auto have id'': "(?mu^n * mu) *s y + ?mu^n *s y = ?mu^Suc n *s y" by (simp add: algebra_simps) show ?case using le unfolding id'' . qed (simp add: matrix_vector_mul_lid) have lt: "0 < cc * y $ i" for i using lt[of i] le[rule_format, of i] by auto have "y $ i > 0" for i using lt[of i] cc by (rule zero_less_mult_pos) thus "lt_vec 0 y" by auto assume **: "mu = sr \ A *v y = mu *s y" { assume "A *v y = mu *s y" with y have "eigen_vector A y mu" unfolding X_def eigen_vector_def by auto hence "eigen_vector cA (map_vector c y) (c mu)" unfolding of_real_hom.eigen_vector_hom . from eigen_vector_norm_sr[OF this] * have "mu = sr" by auto } with ** have mu_sr: "mu = sr" by auto from eq1[folded vector_smult_distrib] have 0: "?w *v (sr *s y - A *v y) = 0" unfolding matrix_vector_right_distrib_diff by simp have le0: "le_vec 0 (sr *s y - A *v y)" using assms(2)[unfolded mu_sr] by auto have "sr *s y - A *v y = 0" using pos_rowvector_mult_0[OF w(2) 0 le0] . hence ev_y: "A *v y = sr *s y" by auto show "mu = sr \ A *v y = mu *s y" using ev_y mu_sr by auto qed lemma nonnegative_eigenvector_has_ev_sr: assumes "eigen_vector A v mu" and le: "le_vec 0 v" shows "mu = sr" proof - from assms(1)[unfolded eigen_vector_def] have v: "v \ 0" and ev: "A *v v = mu *s v" by auto from le v have v: "v \ X" unfolding X_def by auto from ev have "le_vec (A *v v) (mu *s v)" by auto from le_vec_A_mu[OF v this] ev show ?thesis by auto qed lemma similar_matrix_rotation: assumes ev: "eigen_value cA \" and \: "cmod \ = sr" - shows "similar_matrix (cis (arg \) *k cA) cA" + shows "similar_matrix (cis (Arg \) *k cA) cA" proof - from ev obtain y where ev: "eigen_vector cA y \" unfolding eigen_value_def by auto let ?y = "norm_v y" note maps = map_vector_def map_matrix_def define yp where "yp = norm_v y" let ?yp = "map_vector c yp" have yp: "yp \ X" unfolding yp_def by (rule norm_v_X[OF ev]) from ev[unfolded eigen_vector_def] have ev_y: "cA *v y = \ *s y" by auto from ev_le_vec[OF ev, unfolded \, folded yp_def] have 1: "le_vec (sr *s yp) (A *v yp)" by simp from rho_le_x_Ax_imp_rho_le_rx[OF yp 1] have "sr \ r yp" by auto with ev_inequalities[OF ev, folded yp_def] have 2: "r yp = sr" by auto have ev_yp: "A *v yp = sr *s yp" and pos_yp: "lt_vec 0 yp" using sr_imp_eigen_vector_main[OF yp 2] sr_u_pos[OF yp 2] by auto - define D where "D = diagvector (\ j. cis (arg (y $ j)))" - define inv_D where "inv_D = diagvector (\ j. cis (- arg (y $ j)))" + define D where "D = diagvector (\ j. cis (Arg (y $ j)))" + define inv_D where "inv_D = diagvector (\ j. cis (- Arg (y $ j)))" have DD: "inv_D ** D = mat 1" "D ** inv_D = mat 1" unfolding D_def inv_D_def by (auto simp add: diagvector_eq_mat cis_mult) { fix i - have "(D *v ?yp) $ i = cis (arg (y $ i)) * c (cmod (y $ i))" + have "(D *v ?yp) $ i = cis (Arg (y $ i)) * c (cmod (y $ i))" unfolding D_def yp_def by (simp add: maps) also have "\ = y $ i" by (simp add: cis_mult_cmod_id) also note calculation } hence y_D_yp: "y = D *v ?yp" by (auto simp: vec_eq_iff) - define \ where "\ = arg \" + define \ where "\ = Arg \" let ?\ = "cis (- \)" have [simp]: "cis (- \) * rcis sr \ = sr" unfolding cis_rcis_eq rcis_mult by simp - have \: "\ = rcis sr \" unfolding \_def \[symmetric] rcis_cmod_arg .. + have \: "\ = rcis sr \" unfolding \_def \[symmetric] rcis_cmod_Arg .. define F where "F = ?\ *k (inv_D ** cA ** D)" have "cA *v (D *v ?yp) = \ *s y" unfolding y_D_yp[symmetric] ev_y by simp also have "inv_D *v \ = \ *s ?yp" unfolding vector_smult_distrib y_D_yp matrix_vector_mul_assoc DD matrix_vector_mul_lid .. also have "?\ *s \ = sr *s ?yp" unfolding \ by simp also have "\ = map_vector c (sr *s yp)" unfolding vec_eq_iff by (auto simp: maps) also have "\ = cA *v ?yp" unfolding ev_yp[symmetric] by (auto simp: maps matrix_vector_mult_def) finally have F: "F *v ?yp = cA *v ?yp" unfolding F_def matrix_scalar_vector_ac[symmetric] unfolding matrix_vector_mul_assoc[symmetric] vector_smult_distrib . - have prod: "inv_D ** cA ** D = (\ i j. cis (- arg (y $ i)) * cA $ i $ j * cis (arg (y $ j)))" + have prod: "inv_D ** cA ** D = (\ i j. cis (- Arg (y $ i)) * cA $ i $ j * cis (Arg (y $ j)))" unfolding inv_D_def D_def diagvector_mult_right diagvector_mult_left by simp { fix i j - have "cmod (F $ i $ j) = cmod (?\ * cA $h i $h j * (cis (- arg (y $h i)) * cis (arg (y $h j))))" + have "cmod (F $ i $ j) = cmod (?\ * cA $h i $h j * (cis (- Arg (y $h i)) * cis (Arg (y $h j))))" unfolding F_def prod vec_lambda_beta matrix_scalar_mult_def by (simp only: ac_simps) also have "\ = A $ i $ j" unfolding cis_mult unfolding norm_mult by simp also note calculation } hence FA: "map_matrix norm F = A" unfolding maps by auto let ?F = "map_matrix c (map_matrix norm F)" let ?G = "?F - F" let ?Re = "map_matrix Re" from F[folded FA] have 0: "?G *v ?yp = 0" unfolding matrix_diff_vect_distrib by simp have "?Re ?G *v yp = map_vector Re (?G *v ?yp)" unfolding maps matrix_vector_mult_def vec_lambda_beta Re_sum by auto also have "\ = 0" unfolding 0 by (simp add: vec_eq_iff maps) finally have 0: "?Re ?G *v yp = 0" . have "?Re ?G = 0" by (rule pos_matrix_mult_0[OF _ pos_yp 0], auto simp: maps complex_Re_le_cmod) hence "?F = F" by (auto simp: maps vec_eq_iff cmod_eq_Re) with FA have AF: "cA = F" by simp from arg_cong[OF this, of "\ A. cis \ *k A"] have sim: "cis \ *k cA = inv_D ** cA ** D" unfolding F_def matrix.scale_scale cis_mult by simp have "similar_matrix (cis \ *k cA) cA" unfolding similar_matrix_def similar_matrix_wit_def sim by (rule exI[of _ inv_D], rule exI[of _ D], auto simp: DD) thus ?thesis unfolding \_def . qed lemma assumes ev: "eigen_value cA \" and \: "cmod \ = sr" shows maximal_eigen_value_order_1: "order \ (charpoly cA) = 1" - and maximal_eigen_value_rotation: "eigen_value cA (x * cis (arg \)) = eigen_value cA x" - "eigen_value cA (x / cis (arg \)) = eigen_value cA x" + and maximal_eigen_value_rotation: "eigen_value cA (x * cis (Arg \)) = eigen_value cA x" + "eigen_value cA (x / cis (Arg \)) = eigen_value cA x" proof - - let ?a = "cis (arg \)" + let ?a = "cis (Arg \)" let ?p = "charpoly cA" from similar_matrix_rotation[OF ev \] have "similar_matrix (?a *k cA) cA" . from similar_matrix_charpoly[OF this] have id: "charpoly (?a *k cA) = ?p" . have a: "?a \ 0" by simp from order_charpoly_smult[OF this, of _ cA, unfolded id] have order_neg: "order x ?p = order (x / ?a) ?p" for x . have order_pos: "order x ?p = order (x * ?a) ?p" for x using order_neg[symmetric, of "x * ?a"] by simp note order_neg[of \] also have id: "\ / ?a = sr" unfolding \[symmetric] by (metis a cis_mult_cmod_id nonzero_mult_div_cancel_left) also have sr: "order \ ?p = 1" unfolding multiplicity_sr_1[symmetric] charpoly_of_real by (rule map_poly_inj_idom_divide_hom.order_hom, unfold_locales) finally show *: "order \ ?p = 1" . show "eigen_value cA (x * ?a) = eigen_value cA x" using order_pos unfolding eigen_value_root_charpoly order_root by auto show "eigen_value cA (x / ?a) = eigen_value cA x" using order_neg unfolding eigen_value_root_charpoly order_root by auto qed lemma maximal_eigen_values_group: assumes M: "M = {ev :: complex. eigen_value cA ev \ cmod ev = sr}" and a: "rcis sr \ \ M" and b: "rcis sr \ \ M" shows "rcis sr (\ + \) \ M" "rcis sr (\ - \) \ M" "rcis sr 0 \ M" proof - { fix a assume *: "rcis sr a \ M" - have id: "cis (arg (rcis sr a)) = cis a" + have id: "cis (Arg (rcis sr a)) = cis a" by (smt * M mem_Collect_eq nonzero_mult_div_cancel_left of_real_eq_0_iff - rcis_cmod_arg rcis_def sr_pos) + rcis_cmod_Arg rcis_def sr_pos) from *[unfolded assms] have "eigen_value cA (rcis sr a)" "cmod (rcis sr a) = sr" by auto from maximal_eigen_value_rotation[OF this, unfolded id] have "eigen_value cA (x * cis a) = eigen_value cA x" "eigen_value cA (x / cis a) = eigen_value cA x" for x by auto } note * = this from *(1)[OF b, of "rcis sr \"] a show "rcis sr (\ + \) \ M" unfolding M by auto from *(2)[OF a, of "rcis sr \"] a show "rcis sr 0 \ M" unfolding M by auto from *(2)[OF b, of "rcis sr \"] a show "rcis sr (\ - \) \ M" unfolding M by auto qed lemma maximal_eigen_value_roots_of_unity_rotation: assumes M: "M = {ev :: complex. eigen_value cA ev \ cmod ev = sr}" and kM: "k = card M" shows "k \ 0" "k \ CARD('n)" "\ f. charpoly A = (monom 1 k - [:sr^k:]) * f \ (\ x. poly (map_poly c f) x = 0 \ cmod x < sr)" "M = (*) (c sr) ` (\ i. (cis (of_nat i * 2 * pi / k))) ` {0 ..< k}" "M = (*) (c sr) ` { x :: complex. x ^ k = 1}" "(*) (cis (2 * pi / k)) ` Spectrum cA = Spectrum cA" unfolding kM proof - let ?M = "card M" note fin = finite_spectrum[of cA] note char = degree_monic_charpoly[of cA] have "?M \ card (Collect (eigen_value cA))" by (rule card_mono[OF fin], unfold M, auto) also have "Collect (eigen_value cA) = {x. poly (charpoly cA) x = 0}" unfolding eigen_value_root_charpoly by auto also have "card \ \ degree (charpoly cA)" by (rule poly_roots_degree, insert char, auto) also have "\ = CARD('n)" using char by simp finally show "?M \ CARD ('n)" . from finite_subset[OF _ fin, of M] have finM: "finite M" unfolding M by blast from finite_distinct_list[OF this] obtain m where Mm: "M = set m" and dist: "distinct m" by auto from Mm dist have card: "?M = length m" by (auto simp: distinct_card) have sr: "sr \ set m" using eigen_value_sr_c sr_pos unfolding Mm[symmetric] M by auto - define s where "s = sort_key arg m" - define a where "a = map arg s" + define s where "s = sort_key Arg m" + define a where "a = map Arg s" let ?k = "length a" from dist Mm card sr have s: "M = set s" "distinct s" "sr \ set s" and card: "?M = ?k" and sorted: "sorted a" unfolding s_def a_def by auto have map_s: "map ((*) (c sr)) (map cis a) = s" unfolding map_map o_def a_def proof (rule map_idI) fix x assume "x \ set s" from this[folded s(1), unfolded M] have id: "cmod x = sr" by auto - show "sr * cis (arg x) = x" - by (subst (5) rcis_cmod_arg[symmetric], unfold id[symmetric] rcis_def, simp) + show "sr * cis (Arg x) = x" + by (subst (5) rcis_cmod_Arg[symmetric], unfold id[symmetric] rcis_def, simp) qed from s(2)[folded map_s, unfolded distinct_map] have a: "distinct a" "inj_on cis (set a)" by auto from s(3) obtain aa a' where a_split: "a = aa # a'" unfolding a_def by (cases s, auto) - from arg_bounded have bounded: "x \ set a \ - pi < x \ x \ pi" for x unfolding a_def by auto + from Arg_bounded have bounded: "x \ set a \ - pi < x \ x \ pi" for x unfolding a_def by auto from bounded[of aa, unfolded a_split] have aa: "- pi < aa \ aa \ pi" by auto let ?aa = "aa + 2 * pi" define args where "args = a @ [?aa]" let ?diff = "\ i. args ! Suc i - args ! i" have bnd: "x \ set a \ x < ?aa" for x using aa bounded[of x] by auto hence aa_a: "?aa \ set a" by fast have sorted: "sorted args" unfolding args_def using sorted unfolding sorted_append by (insert bnd, auto simp: order.strict_iff_order) have dist: "distinct args" using a aa_a unfolding args_def distinct_append by auto have sum: "(\ i < ?k. ?diff i) = 2 * pi" unfolding sum_lessThan_telescope args_def a_split by simp have k: "?k \ 0" unfolding a_split by auto let ?A = "?diff ` {..< ?k}" let ?Min = "Min ?A" define Min where "Min = ?Min" have "?Min = (?k * ?Min) / ?k" using k by auto also have "?k * ?Min = (\ i < ?k. ?Min)" by auto also have "\ / ?k \ (\ i < ?k. ?diff i) / ?k" by (rule divide_right_mono[OF sum_mono[OF Min_le]], auto) also have "\ = 2 * pi / ?k" unfolding sum .. finally have Min: "Min \ 2 * pi / ?k" unfolding Min_def by auto have lt: "i < ?k \ args ! i < args ! (Suc i)" for i using sorted[unfolded sorted_iff_nth_mono, rule_format, of i "Suc i"] dist[unfolded distinct_conv_nth, rule_format, of "Suc i" i] by (auto simp: args_def) let ?c = "\ i. rcis sr (args ! i)" have hda[simp]: "hd a = aa" unfolding a_split by simp have Min0: "Min > 0" using lt unfolding Min_def by (subst Min_gr_iff, insert k, auto) have Min_A: "Min \ ?A" unfolding Min_def by (rule Min_in, insert k, auto) { fix i :: nat assume i: "i < length args" hence "?c i = rcis sr ((a @ [hd a]) ! i)" by (cases "i = ?k", auto simp: args_def nth_append rcis_def) also have "\ \ set (map (rcis sr) (a @ [hd a]))" using i unfolding args_def set_map unfolding set_conv_nth by auto also have "\ = rcis sr ` set a" unfolding a_split by auto also have "\ = M" unfolding s(1) map_s[symmetric] set_map image_image by (rule image_cong[OF refl], auto simp: rcis_def) finally have "?c i \ M" by auto } note ciM = this { fix i :: nat assume i: "i < ?k" hence "i < length args" "Suc i < length args" unfolding args_def by auto from maximal_eigen_values_group[OF M ciM[OF this(2)] ciM[OF this(1)]] have "rcis sr (?diff i) \ M" by simp } hence Min_M: "rcis sr Min \ M" using Min_A by force have rcisM: "rcis sr (of_nat n * Min) \ M" for n proof (induct n) case 0 show ?case using sr Mm by auto next case (Suc n) have *: "rcis sr (of_nat (Suc n) * Min) = rcis sr (of_nat n * Min) * cis Min" by (simp add: rcis_mult ring_distribs add.commute) from maximal_eigen_values_group(1)[OF M Suc Min_M] show ?case unfolding * by simp qed let ?list = "map (rcis sr) (map (\ i. of_nat i * Min) [0 ..< ?k])" define list where "list = ?list" have len: "length ?list = ?M" unfolding card by simp from sr_pos have sr0: "sr \ 0" by auto { fix i assume i: "i < ?k" hence *: "0 \ real i * Min" using Min0 by auto from i have "real i < real ?k" by auto from mult_strict_right_mono[OF this Min0] have "real i * Min < real ?k * Min" by simp also have "\ \ real ?k * (2 * pi / real ?k)" by (rule mult_left_mono[OF Min], auto) also have "\ = 2 * pi" using k by simp finally have "real i * Min < 2 * pi" . note * this } note prod_pi = this have dist: "distinct ?list" unfolding distinct_map[of "rcis sr"] proof (rule conjI[OF _ inj_on_subset[OF rcis_inj_on[OF sr0]]]) show "distinct (map (\ i. of_nat i * Min) [0 ..< ?k])" using Min0 by (auto simp: distinct_map inj_on_def) show "set (map (\i. real i * Min) [0.. {0..<2 * pi}" using prod_pi by auto qed with len have card': "card (set ?list) = ?M" using distinct_card by fastforce have listM: "set ?list \ M" using rcisM by auto from card_subset_eq[OF finM listM card'] have M_list: "M = set ?list" .. let ?piM = "2 * pi / ?M" { assume "Min \ ?piM" with Min have lt: "Min < 2 * pi / ?k" unfolding card by simp from k have "0 < real ?k" by auto from mult_strict_left_mono[OF lt this] k Min0 have k: "0 \ ?k * Min" "?k * Min < 2 * pi" by auto from rcisM[of ?k, unfolded M_list] have "rcis sr (?k * Min) \ set ?list" by auto then obtain i where i: "i < ?k" and id: "rcis sr (?k * Min) = rcis sr (i * Min)" by auto from inj_onD[OF inj_on_subset[OF rcis_inj_on[OF sr0], of "{?k * Min, i * Min}"] id] prod_pi[OF i] k have "?k * Min = i * Min" by auto with Min0 i have False by auto } hence Min: "Min = ?piM" by auto show cM: "?M \ 0" unfolding card using k by auto let ?f = "(\ i. cis (of_nat i * 2 * pi / ?M))" note M_list also have "set ?list = (*) (c sr) ` (\ i. cis (of_nat i * Min)) ` {0 ..< ?k}" unfolding set_map image_image by (rule image_cong, insert sr_pos, auto simp: rcis_mult rcis_def) finally show M_cis: "M = (*) (c sr) ` ?f ` {0 ..< ?M}" unfolding card Min by (simp add: mult.assoc) thus M_pow: "M = (*) (c sr) ` { x :: complex. x ^ ?M = 1}" using roots_of_unity[OF cM] by simp let ?rphi = "rcis sr (2 * pi / ?M)" let ?phi = "cis (2 * pi / ?M)" from Min_M[unfolded Min] have ev: "eigen_value cA ?rphi" unfolding M by auto have cm: "cmod ?rphi = sr" using sr_pos by simp - have id: "cis (arg ?rphi) = cis (arg ?phi) * cmod ?phi" + have id: "cis (Arg ?rphi) = cis (Arg ?phi) * cmod ?phi" unfolding arg_rcis_cis[OF sr_pos] by simp also have "\ = ?phi" unfolding cis_mult_cmod_id .. - finally have id: "cis (arg ?rphi) = ?phi" . + finally have id: "cis (Arg ?rphi) = ?phi" . define phi where "phi = ?phi" have phi: "phi \ 0" unfolding phi_def by auto note max = maximal_eigen_value_rotation[OF ev cm, unfolded id phi_def[symmetric]] have "((*) phi) ` Spectrum cA = Spectrum cA" (is "?L = ?R") proof - { fix x have *: "x \ ?L \ x \ ?R" for x using max(2)[of x] phi unfolding Spectrum_def by auto moreover { assume "x \ ?R" hence "eigen_value cA x" unfolding Spectrum_def by auto from this[folded max(2)[of x]] have "x / phi \ ?R" unfolding Spectrum_def by auto from imageI[OF this, of "(*) phi"] have "x \ ?L" using phi by auto } note this * } thus ?thesis by blast qed from this[unfolded phi_def] show "(*) (cis (2 * pi / real (card M))) ` Spectrum cA = Spectrum cA" . let ?p = "monom 1 k - [:sr^k:]" let ?cp = "monom 1 k - [:(c sr)^k:]" let ?one = "1 :: complex" let ?list = "map (rcis sr) (map (\ i. of_nat i * ?piM) [0 ..< card M])" interpret c: field_hom c .. interpret p: map_poly_inj_idom_divide_hom c .. have cp: "?cp = map_poly c ?p" by (simp add: hom_distribs) have M_list: "M = set ?list" using M_list[unfolded Min card[symmetric]] . have dist: "distinct ?list" using dist[unfolded Min card[symmetric]] . have k0: "k \ 0" using k[folded card] assms by auto have "?cp = (monom 1 k + (- [:(c sr)^k:]))" by simp also have "degree \ = k" by (subst degree_add_eq_left, insert k0, auto simp: degree_monom_eq) finally have deg: "degree ?cp = k" . from deg k0 have cp0: "?cp \ 0" by auto have "{x. poly ?cp x = 0} = {x. x^k = (c sr)^k}" unfolding poly_diff poly_monom by simp also have "\ \ M" proof - { fix x assume id: "x^k = (c sr)^k" from sr_pos k0 have "(c sr)^k \ 0" by auto with arg_cong[OF id, of "\ x. x / (c sr)^k"] have "(x / c sr)^k = 1" unfolding power_divide by auto hence "c sr * (x / c sr) \ M" by (subst M_pow, unfold kM[symmetric], blast) also have "c sr * (x / c sr) = x" using sr_pos by auto finally have "x \ M" . } thus ?thesis by auto qed finally have cp_M: "{x. poly ?cp x = 0} \ M" . have "k = card (set ?list)" unfolding distinct_card[OF dist] by (simp add: kM) also have "\ \ card {x. poly ?cp x = 0}" proof (rule card_mono[OF poly_roots_finite[OF cp0]]) { fix x assume "x \ set ?list" then obtain i where x: "x = rcis sr (real i * ?piM)" by auto have "x^k = (c sr)^k" unfolding x DeMoivre2 kM by simp (metis mult.assoc of_real_power rcis_times_2pi) hence "poly ?cp x = 0" unfolding poly_diff poly_monom by simp } thus "set ?list \ {x. poly ?cp x = 0}" by auto qed finally have k_card: "k \ card {x. poly ?cp x = 0}" . from k_card cp_M finM have M_id: "M = {x. poly ?cp x = 0}" unfolding kM by (metis card_seteq) have dvdc: "?cp dvd charpoly cA" proof (rule poly_roots_dvd[OF cp0 deg k_card]) from cp_M show "{x. poly ?cp x = 0} \ {x. poly (charpoly cA) x = 0}" unfolding M eigen_value_root_charpoly by auto qed from this[unfolded charpoly_of_real cp p.hom_dvd_iff] have dvd: "?p dvd charpoly A" . from this[unfolded dvd_def] obtain f where decomp: "charpoly A = ?p * f" by blast let ?f = "map_poly c f" have decompc: "charpoly cA = ?cp * ?f" unfolding charpoly_of_real decomp p.hom_mult cp .. show "\ f. charpoly A = (monom 1 ?M - [:sr^?M:]) * f \ (\ x. poly (map_poly c f) x = 0 \ cmod x < sr)" unfolding kM[symmetric] proof (intro exI conjI allI impI, rule decomp) fix x assume f: "poly ?f x = 0" hence ev: "eigen_value cA x" unfolding decompc p.hom_mult eigen_value_root_charpoly by auto hence le: "cmod x \ sr" using eigen_value_norm_sr by auto { assume max: "cmod x = sr" hence "x \ M" unfolding M using ev by auto hence "poly ?cp x = 0" unfolding M_id by auto hence dvd1: "[: -x, 1 :] dvd ?cp" unfolding poly_eq_0_iff_dvd by auto from f[unfolded poly_eq_0_iff_dvd] have dvd2: "[: -x, 1 :] dvd ?f" by auto from char have 0: "charpoly cA \ 0" by auto from mult_dvd_mono[OF dvd1 dvd2] have "[: -x, 1 :]^2 dvd (charpoly cA)" unfolding decompc power2_eq_square . from order_max[OF this 0] maximal_eigen_value_order_1[OF ev max] have False by auto } with le show "cmod x < sr" by argo qed qed lemmas pf_main = eigen_value_sr eigen_vector_z_sr (* sr is eigenvalue *) eigen_value_norm_sr (* it is maximal among all complex eigenvalues *) z_pos (* it's eigenvector is positive *) multiplicity_sr_1 (* the algebr. multiplicity is 1 *) nonnegative_eigenvector_has_ev_sr (* every non-negative real eigenvector has sr as eigenvalue *) maximal_eigen_value_order_1 (* all maximal eigenvalues have order 1 *) maximal_eigen_value_roots_of_unity_rotation (* the maximal eigenvalues are precisely the k-th roots of unity for some k \ dim A *) lemmas pf_main_connect = pf_main(1,3,5,7,8-10)[unfolded sr_spectral_radius] sr_pos[unfolded sr_spectral_radius] end end diff --git a/thys/Perron_Frobenius/Roots_Unity.thy b/thys/Perron_Frobenius/Roots_Unity.thy --- a/thys/Perron_Frobenius/Roots_Unity.thy +++ b/thys/Perron_Frobenius/Roots_Unity.thy @@ -1,585 +1,580 @@ (* author: Thiemann *) section \Roots of Unity\ theory Roots_Unity imports Polynomial_Factorization.Order_Polynomial "HOL-Computational_Algebra.Fundamental_Theorem_Algebra" Polynomial_Interpolation.Ring_Hom_Poly begin -lemma cis_mult_cmod_id: "cis (arg x) * of_real (cmod x) = x" - using rcis_cmod_arg[unfolded rcis_def] by (simp add: ac_simps) +lemma cis_mult_cmod_id: "cis (Arg x) * of_real (cmod x) = x" + using rcis_cmod_Arg[unfolded rcis_def] by (simp add: ac_simps) lemma rcis_mult_cis[simp]: "rcis n a * cis b = rcis n (a + b)" unfolding cis_rcis_eq rcis_mult by simp lemma rcis_div_cis[simp]: "rcis n a / cis b = rcis n (a - b)" unfolding cis_rcis_eq rcis_divide by simp lemma cis_plus_2pi[simp]: "cis (x + 2 * pi) = cis x" by (auto simp: complex_eq_iff) lemma cis_plus_2pi_neq_1: assumes x: "0 < x" "x < 2 * pi" shows "cis x \ 1" proof - from x have "cos x \ 1" by (smt cos_2pi_minus cos_monotone_0_pi cos_zero) thus ?thesis by (auto simp: complex_eq_iff) qed lemma cis_times_2pi[simp]: "cis (of_nat n * 2 * pi) = 1" proof (induct n) case (Suc n) have "of_nat (Suc n) * 2 * pi = of_nat n * 2 * pi + 2 * pi" by (simp add: distrib_right) also have "cis \ = 1" unfolding cis_plus_2pi Suc .. finally show ?case . qed simp -declare cis_pi[simp] - -lemma cis_pi_2[simp]: "cis (pi / 2) = \" - by (auto simp: complex_eq_iff) - lemma cis_add_pi[simp]: "cis (pi + x) = - cis x" by (auto simp: complex_eq_iff) lemma cis_3_pi_2[simp]: "cis (pi * 3 / 2) = - \" proof - have "cis (pi * 3 / 2) = cis (pi + pi / 2)" by (rule arg_cong[of _ _ cis], simp) also have "\ = - \" unfolding cis_add_pi by simp finally show ?thesis . qed lemma rcis_plus_2pi[simp]: "rcis y (x + 2 * pi) = rcis y x" unfolding rcis_def by simp lemma rcis_times_2pi[simp]: "rcis r (of_nat n * 2 * pi) = of_real r" unfolding rcis_def cis_times_2pi by simp -lemma arg_rcis_cis: assumes n: "n > 0" shows "arg (rcis n x) = arg (cis x)" - using arg_bounded arg_unique cis_arg complex_mod_rcis n rcis_def sgn_eq by auto +lemma arg_rcis_cis: assumes n: "n > 0" shows "Arg (rcis n x) = Arg (cis x)" + using Arg_bounded arg_unique cis_Arg complex_mod_rcis n rcis_def sgn_eq by auto -lemma arg_eqD: assumes "arg (cis x) = arg (cis y)" "-pi < x" "x \ pi" "-pi < y" "y \ pi" +lemma arg_eqD: assumes "Arg (cis x) = Arg (cis y)" "-pi < x" "x \ pi" "-pi < y" "y \ pi" shows "x = y" using assms(1) unfolding arg_unique[OF sgn_cis assms(2-3)] arg_unique[OF sgn_cis assms(4-5)] . lemma rcis_inj_on: assumes r: "r \ 0" shows "inj_on (rcis r) {0 ..< 2 * pi}" proof (rule inj_onI, goal_cases) case (1 x y) from arg_cong[OF 1(3), of "\ x. x / r"] have "cis x = cis y" using r by (simp add: rcis_def) from arg_cong[OF this, of "\ x. inverse x"] have "cis (-x) = cis (-y)" by simp from arg_cong[OF this, of uminus] have *: "cis (-x + pi) = cis (-y + pi)" by (auto simp: complex_eq_iff) have "- x + pi = - y + pi" - by (rule arg_eqD[OF arg_cong[OF *, of arg]], insert 1(1-2), auto) + by (rule arg_eqD[OF arg_cong[OF *, of Arg]], insert 1(1-2), auto) thus ?case by simp qed lemma cis_inj_on: "inj_on cis {0 ..< 2 * pi}" using rcis_inj_on[of 1] unfolding rcis_def by auto definition root_unity :: "nat \ 'a :: comm_ring_1 poly" where "root_unity n = monom 1 n - 1" lemma poly_root_unity: "poly (root_unity n) x = 0 \ x^n = 1" unfolding root_unity_def by (simp add: poly_monom) lemma degree_root_unity[simp]: "degree (root_unity n) = n" (is "degree ?p = _") proof - have p: "?p = monom 1 n + (-1)" unfolding root_unity_def by auto show ?thesis proof (cases n) case 0 thus ?thesis unfolding p by simp next case (Suc m) show ?thesis unfolding p unfolding Suc by (subst degree_add_eq_left, auto simp: degree_monom_eq) qed qed lemma zero_root_unit[simp]: "root_unity n = 0 \ n = 0" (is "?p = 0 \ _") proof (cases "n = 0") case True thus ?thesis unfolding root_unity_def by simp next case False from degree_root_unity[of n] False have "degree ?p \ 0" by auto hence "?p \ 0" by fastforce thus ?thesis using False by auto qed definition prod_root_unity :: "nat list \ 'a :: idom poly" where "prod_root_unity ns = prod_list (map root_unity ns)" lemma poly_prod_root_unity: "poly (prod_root_unity ns) x = 0 \ (\k\set ns. x ^ k = 1)" unfolding prod_root_unity_def by (simp add: poly_prod_list prod_list_zero_iff o_def image_def poly_root_unity) lemma degree_prod_root_unity[simp]: "0 \ set ns \ degree (prod_root_unity ns) = sum_list ns" unfolding prod_root_unity_def by (subst degree_prod_list_eq, auto simp: o_def) lemma zero_prod_root_unit[simp]: "prod_root_unity ns = 0 \ 0 \ set ns" unfolding prod_root_unity_def prod_list_zero_iff by auto lemma roots_of_unity: assumes n: "n \ 0" shows "(\ i. (cis (of_nat i * 2 * pi / n))) ` {0 ..< n} = { x :: complex. x ^ n = 1}" (is "?prod = ?Roots") "{x. poly (root_unity n) x = 0} = { x :: complex. x ^ n = 1}" "card { x :: complex. x ^ n = 1} = n" proof (atomize(full), goal_cases) case 1 let ?one = "1 :: complex" let ?p = "monom ?one n - 1" have degM: "degree (monom ?one n) = n" by (rule degree_monom_eq, simp) have "degree ?p = degree (monom ?one n + (-1))" by simp also have "\ = degree (monom ?one n)" by (rule degree_add_eq_left, insert n, simp add: degM) finally have degp: "degree ?p = n" unfolding degM . with n have p: "?p \ 0" by auto have roots: "?Roots = {x. poly ?p x = 0}" unfolding poly_diff poly_monom by simp also have "finite \" by (rule poly_roots_finite[OF p]) finally have fin: "finite ?Roots" . have sub: "?prod \ ?Roots" proof fix x assume "x \ ?prod" then obtain i where x: "x = cis (real i * 2 * pi / n)" by auto have "x ^ n = cis (real i * 2 * pi)" unfolding x DeMoivre using n by simp also have "\ = 1" by simp finally show "x \ ?Roots" by auto qed have Rn: "card ?Roots \ n" unfolding roots by (rule poly_roots_degree[of ?p, unfolded degp, OF p]) have "\ = card {0 ..< n}" by simp also have "\ = card ?prod" proof (rule card_image[symmetric], rule inj_onI, goal_cases) case (1 x y) { fix m assume "m < n" hence "real m < real n" by simp from mult_strict_right_mono[OF this, of "2 * pi / real n"] n have "real m * 2 * pi / real n < real n * 2 * pi / real n" by simp hence "real m * 2 * pi / real n < 2 * pi" using n by simp } note [simp] = this have 0: "(1 :: real) \ 0" using n by auto have "real x * 2 * pi / real n = real y * 2 * pi / real n" by (rule inj_onD[OF rcis_inj_on 1(3)[unfolded cis_rcis_eq]], insert 1(1-2), auto) with n show "x = y" by auto qed finally have cn: "card ?prod = n" .. with Rn have "card ?prod \ card ?Roots" by auto with card_mono[OF fin sub] have card: "card ?prod = card ?Roots" by auto have "?prod = ?Roots" by (rule card_subset_eq[OF fin sub card]) from this roots[symmetric] cn[unfolded this] show ?case unfolding root_unity_def by blast qed lemma poly_roots_dvd: fixes p :: "'a :: field poly" assumes "p \ 0" and "degree p = n" and "card {x. poly p x = 0} \ n" and "{x. poly p x = 0} \ {x. poly q x = 0}" shows "p dvd q" proof - from poly_roots_degree[OF assms(1)] assms(2-3) have "card {x. poly p x = 0} = n" by auto from assms(1-2) this assms(4) show ?thesis proof (induct n arbitrary: p q) case (0 p q) from is_unit_iff_degree[OF 0(1)] 0(2) show ?case by blast next case (Suc n p q) let ?P = "{x. poly p x = 0}" let ?Q = "{x. poly q x = 0}" from Suc(4-5) card_gt_0_iff[of ?P] obtain x where x: "poly p x = 0" "poly q x = 0" and fin: "finite ?P" by auto define r where "r = [:-x, 1:]" from x[unfolded poly_eq_0_iff_dvd r_def[symmetric]] obtain p' q' where p: "p = r * p'" and q: "q = r * q'" unfolding dvd_def by auto from Suc(2) have "degree p = degree r + degree p'" unfolding p by (subst degree_mult_eq, auto) with Suc(3) have deg: "degree p' = n" unfolding r_def by auto from Suc(2) p have p'0: "p' \ 0" by auto let ?P' = "{x. poly p' x = 0}" let ?Q' = "{x. poly q' x = 0}" have P: "?P = insert x ?P'" unfolding p poly_mult unfolding r_def by auto have Q: "?Q = insert x ?Q'" unfolding q poly_mult unfolding r_def by auto { assume "x \ ?P'" hence "?P = ?P'" unfolding P by auto from arg_cong[OF this, of card, unfolded Suc(4)] deg have False using poly_roots_degree[OF p'0] by auto } note xp' = this hence xP': "x \ ?P'" by auto have "card ?P = Suc (card ?P')" unfolding P by (rule card_insert_disjoint[OF _ xP'], insert fin[unfolded P], auto) with Suc(4) have card: "card ?P' = n" by auto from Suc(5)[unfolded P Q] xP' have "?P' \ ?Q'" by auto from Suc(1)[OF p'0 deg card this] have IH: "p' dvd q'" . show ?case unfolding p q using IH by simp qed qed lemma root_unity_decomp: assumes n: "n \ 0" shows "root_unity n = prod_list (map (\ i. [:-cis (of_nat i * 2 * pi / n), 1:]) [0 ..< n])" (is "?u = ?p") proof - have deg: "degree ?u = n" by simp note main = roots_of_unity[OF n] have dvd: "?u dvd ?p" proof (rule poly_roots_dvd[OF _ deg]) show "n \ card {x. poly ?u x = 0}" using main by auto show "?u \ 0" using n by auto show "{x. poly ?u x = 0} \ {x. poly ?p x = 0}" unfolding main(2) main(1)[symmetric] poly_prod_list prod_list_zero_iff by auto qed have deg': "degree ?p = n" by (subst degree_prod_list_eq, auto simp: o_def sum_list_triv) have mon: "monic ?u" using deg unfolding root_unity_def using n by auto have mon': "monic ?p" by (rule monic_prod_list, auto) from dvd[unfolded dvd_def] obtain f where puf: "?p = ?u * f" by auto have "degree ?p = degree ?u + degree f" using mon' n unfolding puf by (subst degree_mult_eq, auto) with deg deg' have "degree f = 0" by auto from degree0_coeffs[OF this] obtain a where f: "f = [:a:]" by blast from arg_cong[OF puf, of lead_coeff] mon mon' have "a = 1" unfolding puf f by (cases "a = 0", auto) with f have f: "f = 1" by auto with puf show ?thesis by auto qed lemma order_monic_linear: "order x [:y,1:] = (if y + x = 0 then 1 else 0)" proof (cases "y + x = 0") case True hence "poly [:y,1:] x = 0" by simp from this[unfolded order_root] have "order x [:y,1:] \ 0" by auto moreover from order_degree[of "[:y,1:]" x] have "order x [:y,1:] \ 1" by auto ultimately show ?thesis unfolding True by auto next case False hence "poly [:y,1:] x \ 0" by auto from order_0I[OF this] False show ?thesis by auto qed lemma order_root_unity: fixes x :: complex assumes n: "n \ 0" shows "order x (root_unity n) = (if x^n = 1 then 1 else 0)" (is "order _ ?u = _") proof (cases "x^n = 1") case False with roots_of_unity(2)[OF n] have "poly ?u x \ 0" by auto from False order_0I[OF this] show ?thesis by auto next case True let ?phi = "\ i :: nat. i * 2 * pi / n" from True roots_of_unity(1)[OF n] obtain i where i: "i < n" and x: "x = cis (?phi i)" by force from i have n_split: "[0 ..< n] = [0 ..< i] @ i # [Suc i ..< n]" by (metis le_Suc_ex less_imp_le_nat not_le_imp_less not_less0 upt_add_eq_append upt_conv_Cons) { fix j assume j: "j < n \ j < i" and eq: "cis (?phi i) = cis (?phi j)" from inj_onD[OF cis_inj_on eq] i j n have "i = j" by (auto simp: field_simps) } note inj = this have "order x ?u = 1" unfolding root_unity_decomp[OF n] unfolding x n_split using inj by (subst order_prod_list, force, fastforce simp: order_monic_linear) with True show ?thesis by auto qed lemma order_prod_root_unity: assumes 0: "0 \ set ks" shows "order (x :: complex) (prod_root_unity ks) = length (filter (\ k. x^k = 1) ks)" proof - have "order x (prod_root_unity ks) = (\k\ks. order x (root_unity k))" unfolding prod_root_unity_def by (subst order_prod_list, insert 0, auto simp: o_def) also have "\ = (\k\ks. (if x^k = 1 then 1 else 0))" by (rule arg_cong, rule map_cong, insert 0, force, intro order_root_unity, metis) also have "\ = length (filter (\ k. x^k = 1) ks)" by (subst sum_list_map_filter'[symmetric], simp add: sum_list_triv) finally show ?thesis . qed lemma root_unity_witness: fixes xs :: "complex list" assumes "prod_list (map (\ x. [:-x,1:]) xs) = monom 1 n - 1" shows "x^n = 1 \ x \ set xs" proof - from assms have n0: "n \ 0" by (cases "n = 0", auto simp: prod_list_zero_iff) have "x \ set xs \ poly (prod_list (map (\ x. [:-x,1:]) xs)) x = 0" unfolding poly_prod_list prod_list_zero_iff by auto also have "\ \ x^n = 1" using roots_of_unity(2)[OF n0] unfolding assms root_unity_def by auto finally show ?thesis by auto qed lemma root_unity_explicit: fixes x :: complex shows "(x ^ 1 = 1) \ x = 1" "(x ^ 2 = 1) \ (x \ {1, -1})" "(x ^ 3 = 1) \ (x \ {1, Complex (-1/2) (sqrt 3 / 2), Complex (-1/2) (- sqrt 3 / 2)})" "(x ^ 4 = 1) \ (x \ {1, -1, \, - \})" proof - show "(x ^ 1 = 1) \ x = 1" by (subst root_unity_witness[of "[1]"], code_simp, auto) show "(x ^ 2 = 1) \ (x \ {1, -1})" by (subst root_unity_witness[of "[1,-1]"], code_simp, auto) show "(x ^ 4 = 1) \ (x \ {1, -1, \, - \})" by (subst root_unity_witness[of "[1,-1, \, - \]"], code_simp, auto) have 3: "3 = Suc (Suc (Suc 0))" "1 = [:1:]" by auto show "(x ^ 3 = 1) \ (x \ {1, Complex (-1/2) (sqrt 3 / 2), Complex (-1/2) (- sqrt 3 / 2)})" by (subst root_unity_witness[of "[1, Complex (-1/2) (sqrt 3 / 2), Complex (-1/2) (- sqrt 3 / 2)]"], auto simp: 3 monom_altdef complex_mult complex_eq_iff) qed definition primitive_root_unity :: "nat \ 'a :: power \ bool" where "primitive_root_unity k x = (k \ 0 \ x^k = 1 \ (\ k' < k. k' \ 0 \ x^k' \ 1))" lemma primitive_root_unityD: assumes "primitive_root_unity k x" shows "k \ 0" "x^k = 1" "k' \ 0 \ x^k' = 1 \ k \ k'" proof - note * = assms[unfolded primitive_root_unity_def] from * have **: "k' < k \ k' \ 0 \ x ^ k' \ 1" by auto show "k \ 0" "x^k = 1" using * by auto show "k' \ 0 \ x^k' = 1 \ k \ k'" using ** by force qed lemma primitive_root_unity_exists: assumes "k \ 0" "x ^ k = 1" shows "\ k'. k' \ k \ primitive_root_unity k' x" proof - let ?P = "\ k. x ^ k = 1 \ k \ 0" define k' where "k' = (LEAST k. ?P k)" from assms have Pk: "\ k. ?P k" by auto from LeastI_ex[OF Pk, folded k'_def] have "k' \ 0" "x ^ k' = 1" by auto with not_less_Least[of _ ?P, folded k'_def] have "primitive_root_unity k' x" unfolding primitive_root_unity_def by auto with primitive_root_unityD(3)[OF this assms] show ?thesis by auto qed lemma primitive_root_unity_dvd: fixes x :: "complex" assumes k: "primitive_root_unity k x" shows "x ^ n = 1 \ k dvd n" proof assume "k dvd n" then obtain j where n: "n = k * j" unfolding dvd_def by auto have "x ^ n = (x ^ k) ^ j" unfolding n power_mult by simp also have "\ = 1" unfolding primitive_root_unityD[OF k] by simp finally show "x ^ n = 1" . next assume n: "x ^ n = 1" note k = primitive_root_unityD[OF k] show "k dvd n" proof (cases "n = 0") case n0: False from k(3)[OF n0] n have nk: "n \ k" by force from roots_of_unity[OF k(1)] k(2) obtain i :: nat where xk: "x = cis (i * 2 * pi / k)" and ik: "i < k" by force from roots_of_unity[OF n0] n obtain j :: nat where xn: "x = cis (j * 2 * pi / n)" and jn: "j < n" by force have cop: "coprime i k" proof (rule gcd_eq_1_imp_coprime) from k(1) have "gcd i k \ 0" by auto from gcd_coprime_exists[OF this] this obtain i' k' g where *: "i = i' * g" "k = k' * g" "g \ 0" and g: "g = gcd i k" by blast from *(2) k(1) have k': "k' \ 0" by auto have "x = cis (i * 2 * pi / k)" by fact also have "i * 2 * pi / k = i' * 2 * pi / k'" unfolding * using *(3) by auto finally have "x ^ k' = 1" by (simp add: DeMoivre k') with k(3)[OF k'] have "k' \ k" by linarith moreover with * k(1) have "g = 1" by auto then show "gcd i k = 1" by (simp add: g) qed from inj_onD[OF cis_inj_on xk[unfolded xn]] n0 k(1) ik jn have "j * real k = i * real n" by (auto simp: field_simps) hence "real (j * k) = real (i * n)" by simp hence eq: "j * k = i * n" by linarith with cop show "k dvd n" by (metis coprime_commute coprime_dvd_mult_right_iff dvd_triv_right) qed auto qed lemma primitive_root_unity_simple_computation: "primitive_root_unity k x = (if k = 0 then False else x ^ k = 1 \ (\ i \ {1 ..< k}. x ^ i \ 1))" unfolding primitive_root_unity_def by auto lemma primitive_root_unity_explicit: fixes x :: complex shows "primitive_root_unity 1 x \ x = 1" "primitive_root_unity 2 x \ x = -1" "primitive_root_unity 3 x \ (x \ {Complex (-1/2) (sqrt 3 / 2), Complex (-1/2) (- sqrt 3 / 2)})" "primitive_root_unity 4 x \ (x \ {\, - \})" proof (atomize(full), goal_cases) case 1 { fix P :: "nat \ bool" have *: "{1 ..< 2 :: nat} = {1}" "{1 ..< 3 :: nat} = {1,2}" "{1 ..< 4 :: nat} = {1,2,3}" by code_simp+ have "(\i\ {1 ..< 2}. P i) = P 1" "(\i\ {1 ..< 3}. P i) \ P 1 \ P 2" "(\i\ {1 ..< 4}. P i) \ P 1 \ P 2 \ P 3" unfolding * by auto } note * = this show ?case unfolding primitive_root_unity_simple_computation root_unity_explicit * by (auto simp: complex_eq_iff) qed function decompose_prod_root_unity_main :: "'a :: field poly \ nat \ nat list \ 'a poly" where "decompose_prod_root_unity_main p k = ( if k = 0 then ([], p) else let q = root_unity k in if q dvd p then if p = 0 then ([],0) else map_prod (Cons k) id (decompose_prod_root_unity_main (p div q) k) else decompose_prod_root_unity_main p (k - 1))" by pat_completeness auto termination by (relation "measure (\ (p,k). degree p + k)", auto simp: degree_div_less) declare decompose_prod_root_unity_main.simps[simp del] lemma decompose_prod_root_unity_main: fixes p :: "complex poly" assumes p: "p = prod_root_unity ks * f" and d: "decompose_prod_root_unity_main p k = (ks',g)" and f: "\ x. cmod x = 1 \ poly f x \ 0" and k: "\ k'. k' > k \ \ root_unity k' dvd p" shows "p = prod_root_unity ks' * f \ f = g \ set ks = set ks'" using d p k proof (induct p k arbitrary: ks ks' rule: decompose_prod_root_unity_main.induct) case (1 p k ks ks') note p = 1(4) note k = 1(5) from k[of "Suc k"] have p0: "p \ 0" by auto hence "p = 0 \ False" by auto note d = 1(3)[unfolded decompose_prod_root_unity_main.simps[of p k] this if_False Let_def] from p0[unfolded p] have ks0: "0 \ set ks" by simp from f[of 1] have f0: "f \ 0" by auto note IH = 1(1)[OF _ refl _ p0] 1(2)[OF _ refl] show ?case proof (cases "k = 0") case True with p k[unfolded this, of "hd ks"] p0 have "ks = []" by (cases ks, auto simp: prod_root_unity_def) with d p True show ?thesis by (auto simp: prod_root_unity_def) next case k0: False note IH = IH[OF k0] from k0 have "k = 0 \ False" by auto note d = d[unfolded this if_False] let ?u = "root_unity k :: complex poly" show ?thesis proof (cases "?u dvd p") case True note IH = IH(1)[OF True] let ?call = "decompose_prod_root_unity_main (p div ?u) k" from True d obtain Ks where rec: "?call = (Ks,g)" and ks': "ks' = (k # Ks)" by (cases ?call, auto) from True have "?u dvd p \ True" by simp note d = d[unfolded this if_True rec] let ?x = "cis (2 * pi / k)" have rt: "poly ?u ?x = 0" unfolding poly_root_unity using cis_times_2pi[of 1] by (simp add: DeMoivre) with True have "poly p ?x = 0" unfolding dvd_def by auto from this[unfolded p] f[of ?x] rt have "poly (prod_root_unity ks) ?x = 0" unfolding poly_root_unity by auto from this[unfolded poly_prod_root_unity] ks0 obtain k' where k': "k' \ set ks" and rt: "?x ^ k' = 1" and k'0: "k' \ 0" by auto let ?u' = "root_unity k' :: complex poly" from k' rt k'0 have rtk': "poly ?u' ?x = 0" unfolding poly_root_unity by auto { let ?phi = " k' * (2 * pi / k)" assume "k' < k" hence "0 < ?phi" "?phi < 2 * pi" using k0 k'0 by (auto simp: field_simps) from cis_plus_2pi_neq_1[OF this] rtk' have False unfolding poly_root_unity DeMoivre .. } hence kk': "k \ k'" by presburger { assume "k' > k" from k[OF this, unfolded p] have "\ ?u' dvd prod_root_unity ks" using dvd_mult2 by auto with k' have False unfolding prod_root_unity_def using prod_list_dvd[of ?u' "map root_unity ks"] by auto } with kk' have kk': "k' = k" by presburger with k' have "k \ set ks" by auto from split_list[OF this] obtain ks1 ks2 where ks: "ks = ks1 @ k # ks2" by auto hence "p div ?u = (?u * (prod_root_unity (ks1 @ ks2) * f)) div ?u" by (simp add: ac_simps p prod_root_unity_def) also have "\ = prod_root_unity (ks1 @ ks2) * f" by (rule nonzero_mult_div_cancel_left, insert k0, auto) finally have id: "p div ?u = prod_root_unity (ks1 @ ks2) * f" . from d have ks': "ks' = k # Ks" by auto have "k < k' \ \ root_unity k' dvd p div ?u" for k' using k[of k'] True by (metis dvd_div_mult_self dvd_mult2) from IH[OF rec id this] have id: "p div root_unity k = prod_root_unity Ks * f" and *: "f = g \ set (ks1 @ ks2) = set Ks" by auto from arg_cong[OF id, of "\ x. x * ?u"] True have "p = prod_root_unity Ks * f * root_unity k" by auto thus ?thesis using * unfolding ks ks' by (auto simp: prod_root_unity_def) next case False from d False have "decompose_prod_root_unity_main p (k - 1) = (ks',g)" by auto note IH = IH(2)[OF False this p] have k: "k - 1 < k' \ \ root_unity k' dvd p" for k' using False k[of k'] k0 by (cases "k' = k", auto) show ?thesis by (rule IH, insert False k, auto) qed qed qed definition "decompose_prod_root_unity p = decompose_prod_root_unity_main p (degree p)" lemma decompose_prod_root_unity: fixes p :: "complex poly" assumes p: "p = prod_root_unity ks * f" and d: "decompose_prod_root_unity p = (ks',g)" and f: "\ x. cmod x = 1 \ poly f x \ 0" and p0: "p \ 0" shows "p = prod_root_unity ks' * f \ f = g \ set ks = set ks'" proof (rule decompose_prod_root_unity_main[OF p d[unfolded decompose_prod_root_unity_def] f]) fix k assume deg: "degree p < k" hence "degree p < degree (root_unity k)" by simp with p0 show "\ root_unity k dvd p" by (simp add: poly_divides_conv0) qed lemma (in comm_ring_hom) hom_root_unity: "map_poly hom (root_unity n) = root_unity n" proof - interpret p: map_poly_comm_ring_hom hom .. show ?thesis unfolding root_unity_def by (simp add: hom_distribs) qed lemma (in idom_hom) hom_prod_root_unity: "map_poly hom (prod_root_unity n) = prod_root_unity n" proof - interpret p: map_poly_comm_ring_hom hom .. show ?thesis unfolding prod_root_unity_def p.hom_prod_list map_map o_def hom_root_unity .. qed lemma (in field_hom) hom_decompose_prod_root_unity_main: "decompose_prod_root_unity_main (map_poly hom p) k = map_prod id (map_poly hom) (decompose_prod_root_unity_main p k)" proof (induct p k rule: decompose_prod_root_unity_main.induct) case (1 p k) let ?h = "map_poly hom" let ?p = "?h p" let ?u = "root_unity k :: 'a poly" let ?u' = "root_unity k :: 'b poly" interpret p: map_poly_inj_idom_divide_hom hom .. have u': "?u' = ?h ?u" unfolding hom_root_unity .. note simp = decompose_prod_root_unity_main.simps let ?rec1 = "decompose_prod_root_unity_main (p div ?u) k" have 0: "?p = 0 \ p = 0" by simp show ?case unfolding simp[of ?p k] simp[of p k] if_distrib[of "map_prod id ?h"] Let_def u' unfolding 0 p.hom_div[symmetric] p.hom_dvd_iff by (rule if_cong[OF refl], force, rule if_cong[OF refl if_cong[OF refl]], force, (subst 1(1), auto, cases ?rec1, auto)[1], (subst 1(2), auto)) qed lemma (in field_hom) hom_decompose_prod_root_unity: "decompose_prod_root_unity (map_poly hom p) = map_prod id (map_poly hom) (decompose_prod_root_unity p)" unfolding decompose_prod_root_unity_def by (subst hom_decompose_prod_root_unity_main, simp) end diff --git a/thys/Poincare_Disc/Poincare_Between.thy b/thys/Poincare_Disc/Poincare_Between.thy --- a/thys/Poincare_Disc/Poincare_Between.thy +++ b/thys/Poincare_Disc/Poincare_Between.thy @@ -1,1257 +1,1257 @@ theory Poincare_Between imports Poincare_Distance begin (* ------------------------------------------------------------------ *) section\H-betweenness in the Poincar\'e model\ (* ------------------------------------------------------------------ *) subsection \H-betwenness expressed by a cross-ratio\ text\The point $v$ is h-between $u$ and $w$ if the cross-ratio between the pairs $u$ and $w$ and $v$ and inverse of $v$ is real and negative.\ definition poincare_between :: "complex_homo \ complex_homo \ complex_homo \ bool" where "poincare_between u v w \ u = v \ v = w \ (let cr = cross_ratio u v w (inversion v) in is_real (to_complex cr) \ Re (to_complex cr) < 0)" subsubsection \H-betwenness is preserved by h-isometries\ text \Since they preserve cross-ratio and inversion, h-isometries (unit disc preserving Möbius transformations and conjugation) preserve h-betweeness.\ lemma unit_disc_fix_moebius_preserve_poincare_between [simp]: assumes "unit_disc_fix M" and "u \ unit_disc" and "v \ unit_disc" and "w \ unit_disc" shows "poincare_between (moebius_pt M u) (moebius_pt M v) (moebius_pt M w) \ poincare_between u v w" proof (cases "u = v \ v = w") case True thus ?thesis using assms unfolding poincare_between_def by auto next case False moreover hence "moebius_pt M u \ moebius_pt M v \ moebius_pt M v \ moebius_pt M w" by auto moreover have "v \ inversion v" "w \ inversion v" using inversion_noteq_unit_disc[of v w] using inversion_noteq_unit_disc[of v v] using \v \ unit_disc\ \w \ unit_disc\ by auto ultimately show ?thesis using assms using unit_circle_fix_moebius_pt_inversion[of M v, symmetric] unfolding poincare_between_def by (simp del: unit_circle_fix_moebius_pt_inversion) qed lemma conjugate_preserve_poincare_between [simp]: assumes "u \ unit_disc" and "v \ unit_disc" and "w \ unit_disc" shows "poincare_between (conjugate u) (conjugate v) (conjugate w) \ poincare_between u v w" proof (cases "u = v \ v = w") case True thus ?thesis using assms unfolding poincare_between_def by auto next case False moreover hence "conjugate u \ conjugate v \ conjugate v \ conjugate w" using conjugate_inj by blast moreover have "v \ inversion v" "w \ inversion v" using inversion_noteq_unit_disc[of v w] using inversion_noteq_unit_disc[of v v] using \v \ unit_disc\ \w \ unit_disc\ by auto ultimately show ?thesis using assms using conjugate_cross_ratio[of v w "inversion v" u] unfolding poincare_between_def by (metis conjugate_id_iff conjugate_involution inversion_def inversion_sym o_apply) qed subsubsection \Some elementary properties of h-betwenness\ lemma poincare_between_nonstrict [simp]: shows "poincare_between u u v" and "poincare_between u v v" by (simp_all add: poincare_between_def) lemma poincare_between_sandwich: assumes "u \ unit_disc" and "v \ unit_disc" assumes "poincare_between u v u" shows "u = v" proof (rule ccontr) assume "\ ?thesis" thus False using assms using inversion_noteq_unit_disc[of v u] using cross_ratio_1[of v u "inversion v"] unfolding poincare_between_def Let_def by auto qed lemma poincare_between_rev: assumes "u \ unit_disc" and "v \ unit_disc" and "w \ unit_disc" shows "poincare_between u v w \ poincare_between w v u" using assms using inversion_noteq_unit_disc[of v w] using inversion_noteq_unit_disc[of v u] using cross_ratio_commute_13[of u v w "inversion v"] using cross_ratio_not_inf[of w "inversion v" v u] using cross_ratio_not_zero[of w v u "inversion v"] using inf_or_of_complex[of "cross_ratio w v u (inversion v)"] unfolding poincare_between_def by (auto simp add: Let_def Im_complex_div_eq_0 Re_divide divide_less_0_iff) subsubsection \H-betwenness and h-collinearity\ text\Three points can be in an h-between relation only when they are h-collinear.\ lemma poincare_between_poincare_collinear [simp]: assumes in_disc: "u \ unit_disc" "v \ unit_disc" "w \ unit_disc" assumes betw: "poincare_between u v w" shows "poincare_collinear {u, v, w}" proof (cases "u = v \ v = w") case True thus ?thesis using assms by auto next case False hence distinct: "distinct [u, v, w, inversion v]" using in_disc inversion_noteq_unit_disc[of v v] inversion_noteq_unit_disc[of v u] inversion_noteq_unit_disc[of v w] using betw poincare_between_sandwich[of w v] by (auto simp add: poincare_between_def Let_def) then obtain H where *: "{u, v, w, inversion v} \ circline_set H" using assms unfolding poincare_between_def using four_points_on_circline_iff_cross_ratio_real[of u v w "inversion v"] by auto hence "H = poincare_line u v" using assms distinct using unique_circline_set[of u v "inversion v"] using poincare_line[of u v] poincare_line_inversion[of u v] unfolding circline_set_def by auto thus ?thesis using * assms False unfolding poincare_collinear_def by (rule_tac x="poincare_line u v" in exI) simp qed lemma poincare_between_poincare_line_uvz: assumes "u \ v" and "u \ unit_disc" and "v \ unit_disc" and "z \ unit_disc" and "poincare_between u v z" shows "z \ circline_set (poincare_line u v)" using assms using poincare_between_poincare_collinear[of u v z] using unique_poincare_line[OF assms(1-3)] unfolding poincare_collinear_def by auto lemma poincare_between_poincare_line_uzv: assumes "u \ v" and "u \ unit_disc" and "v \ unit_disc" and "z \ unit_disc" "poincare_between u z v" shows "z \ circline_set (poincare_line u v)" using assms using poincare_between_poincare_collinear[of u z v] using unique_poincare_line[OF assms(1-3)] unfolding poincare_collinear_def by auto subsubsection \H-betweeness on Euclidean segments\ text\If the three points lie on an h-line that is a Euclidean line (e.g., if it contains zero), h-betweenness can be characterized much simpler than in the definition.\ lemma poincare_between_x_axis_u0v: assumes "is_real u'" and "u' \ 0" and "v' \ 0" shows "poincare_between (of_complex u') 0\<^sub>h (of_complex v') \ is_real v' \ Re u' * Re v' < 0" proof- have "Re u' \ 0" using \is_real u'\ \u' \ 0\ using complex_eq_if_Re_eq by auto have nz: "of_complex u' \ 0\<^sub>h" "of_complex v' \ 0\<^sub>h" by (simp_all add: \u' \ 0\ \v' \ 0\) hence "0\<^sub>h \ of_complex v'" by metis let ?cr = "cross_ratio (of_complex u') 0\<^sub>h (of_complex v') \\<^sub>h" have "is_real (to_complex ?cr) \ Re (to_complex ?cr) < 0 \ is_real v' \ Re u' * Re v' < 0" using cross_ratio_0inf[of v' u'] \v' \ 0\ \u' \ 0\ \is_real u'\ by (metis Re_complex_div_lt_0 Re_mult_real complex_cnj_divide divide_cancel_left eq_cnj_iff_real to_complex_of_complex) thus ?thesis unfolding poincare_between_def inversion_zero using \of_complex u' \ 0\<^sub>h\ \0\<^sub>h \ of_complex v'\ by simp qed lemma poincare_between_u0v: assumes "u \ unit_disc" and "v \ unit_disc" and "u \ 0\<^sub>h" and "v \ 0\<^sub>h" shows "poincare_between u 0\<^sub>h v \ (\ k < 0. to_complex u = cor k * to_complex v)" (is "?P u v") proof (cases "u = v") case True thus ?thesis using assms using inf_or_of_complex[of v] using poincare_between_sandwich[of u "0\<^sub>h"] by auto next case False have "\ u. u \ unit_disc \ u \ 0\<^sub>h \ ?P u v" (is "?P' v") proof (rule wlog_rotation_to_positive_x_axis) fix \ v let ?M = "moebius_pt (moebius_rotation \)" assume 1: "v \ unit_disc" "v \ 0\<^sub>h" assume 2: "?P' (?M v)" show "?P' v" proof (rule allI, rule impI, (erule conjE)+) fix u assume 3: "u \ unit_disc" "u \ 0\<^sub>h" have "poincare_between (?M u) 0\<^sub>h (?M v) \ poincare_between u 0\<^sub>h v" using \u \ unit_disc\ \v \ unit_disc\ using unit_disc_fix_moebius_preserve_poincare_between unit_disc_fix_rotation zero_in_unit_disc by fastforce thus "?P u v" using 1 2[rule_format, of "?M u"] 3 using inf_or_of_complex[of u] inf_or_of_complex[of v] by auto qed next fix x assume 1: "is_real x" "0 < Re x" "Re x < 1" hence "x \ 0" by auto show "?P' (of_complex x)" proof (rule allI, rule impI, (erule conjE)+) fix u assume 2: "u \ unit_disc" "u \ 0\<^sub>h" then obtain u' where "u = of_complex u'" using inf_or_of_complex[of u] by auto show "?P u (of_complex x)" using 1 2 \x \ 0\ \u = of_complex u'\ using poincare_between_rev[of u "0\<^sub>h" "of_complex x"] using poincare_between_x_axis_u0v[of x u'] \is_real x\ apply (auto simp add: cmod_eq_Re) apply (rule_tac x="Re u' / Re x" in exI, simp add: divide_neg_pos algebra_split_simps) using mult_neg_pos mult_pos_neg by blast qed qed fact+ thus ?thesis using assms by auto qed lemma poincare_between_u0v_polar_form: assumes "x \ unit_disc" and "y \ unit_disc" and "x \ 0\<^sub>h" and "y \ 0\<^sub>h" and "to_complex x = cor rx * cis \" "to_complex y = cor ry * cis \" shows "poincare_between x 0\<^sub>h y \ rx * ry < 0" (is "?P x y rx ry") proof- from assms have "rx \ 0" "ry \ 0" using inf_or_of_complex[of x] inf_or_of_complex[of y] by auto have "(\k<0. cor rx = cor k * cor ry ) = (rx * ry < 0)" proof assume "\k<0. cor rx = cor k * cor ry" then obtain k where "k < 0" "cor rx = cor k * cor ry" by auto hence "rx = k * ry" using of_real_eq_iff by fastforce thus "rx * ry < 0" using \k < 0\ \rx \ 0\ \ry \ 0\ by (smt divisors_zero mult_nonneg_nonpos mult_nonpos_nonpos zero_less_mult_pos2) next assume "rx * ry < 0" hence "rx = (rx/ry)*ry" "rx / ry < 0" using \rx \ 0\ \ry \ 0\ by (auto simp add: divide_less_0_iff algebra_split_simps) thus "\k<0. cor rx = cor k * cor ry" using \rx \ 0\ \ry \ 0\ by (rule_tac x="rx / ry" in exI, simp) qed thus ?thesis using assms using poincare_between_u0v[OF assms(1-4)] by auto qed lemma poincare_between_x_axis_0uv: fixes x y :: real assumes "-1 < x" and "x < 1" and "x \ 0" assumes "-1 < y" and "y < 1" and "y \ 0" shows "poincare_between 0\<^sub>h (of_complex x) (of_complex y) \ (x < 0 \ y < 0 \ y \ x) \ (x > 0 \ y > 0 \ x \ y)" (is "?lhs \ ?rhs") proof (cases "x = y") case True thus ?thesis using assms unfolding poincare_between_def by auto next case False let ?x = "of_complex x" and ?y = "of_complex y" have "?x \ unit_disc" "?y \ unit_disc" using assms by auto have distinct: "distinct [0\<^sub>h, ?x, ?y, inversion ?x]" using \x \ 0\ \y \ 0\ \x \ y\ \?x \ unit_disc\ \?y \ unit_disc\ using inversion_noteq_unit_disc[of ?x ?y] using inversion_noteq_unit_disc[of ?x ?x] using inversion_noteq_unit_disc[of ?x "0\<^sub>h"] using of_complex_inj[of x y] by (metis distinct_length_2_or_more distinct_singleton of_complex_zero_iff of_real_eq_0_iff of_real_eq_iff zero_in_unit_disc) let ?cr = "cross_ratio 0\<^sub>h ?x ?y (inversion ?x)" have "Re (to_complex ?cr) = x\<^sup>2 * (x*y - 1) / (x * (y - x))" using \x \ 0\ \x \ y\ unfolding inversion_def by simp (transfer, transfer, auto simp add: vec_cnj_def power2_eq_square field_simps split: if_split_asm) moreover { fix a b :: real assume "b \ 0" hence "a < 0 \ b\<^sup>2 * a < (0::real)" by (metis mult.commute mult_eq_0_iff mult_neg_pos mult_pos_pos not_less_iff_gr_or_eq not_real_square_gt_zero power2_eq_square) } hence "x\<^sup>2 * (x*y - 1) < 0" using assms by (smt minus_mult_minus mult_le_cancel_left1) moreover have "x * (y - x) > 0 \ ?rhs" using \x \ 0\ \y \ 0\ \x \ y\ by (smt mult_le_0_iff) ultimately have *: "Re (to_complex ?cr) < 0 \ ?rhs" by (simp add: divide_less_0_iff) show ?thesis proof assume ?lhs have "is_real (to_complex ?cr)" "Re (to_complex ?cr) < 0" using \?lhs\ distinct unfolding poincare_between_def Let_def by auto thus ?rhs using * by simp next assume ?rhs hence "Re (to_complex ?cr) < 0" using * by simp moreover have "{0\<^sub>h, of_complex (cor x), of_complex (cor y), inversion (of_complex (cor x))} \ circline_set x_axis" using \x \ 0\ is_real_inversion[of "cor x"] using inf_or_of_complex[of "inversion ?x"] by (auto simp del: inversion_of_complex) hence "is_real (to_complex ?cr)" using four_points_on_circline_iff_cross_ratio_real[OF distinct] by auto ultimately show ?lhs using distinct unfolding poincare_between_def Let_def by auto qed qed lemma poincare_between_0uv: assumes "u \ unit_disc" and "v \ unit_disc" and "u \ 0\<^sub>h" and "v \ 0\<^sub>h" shows "poincare_between 0\<^sub>h u v \ (let u' = to_complex u; v' = to_complex v in arg u' = arg v' \ cmod u' \ cmod v')" (is "?P u v") proof (cases "u = v") case True thus ?thesis by simp next case False have "\ v. v \ unit_disc \ v \ 0\<^sub>h \ v \ u \ (poincare_between 0\<^sub>h u v \ (let u' = to_complex u; v' = to_complex v in arg u' = arg v' \ cmod u' \ cmod v'))" (is "?P' u") proof (rule wlog_rotation_to_positive_x_axis) show "u \ unit_disc" "u \ 0\<^sub>h" by fact+ next fix x assume *: "is_real x" "0 < Re x" "Re x < 1" hence "of_complex x \ unit_disc" "of_complex x \ 0\<^sub>h" "of_complex x \ circline_set x_axis" unfolding circline_set_x_axis by (auto simp add: cmod_eq_Re) show "?P' (of_complex x)" proof safe fix v assume "v \ unit_disc" "v \ 0\<^sub>h" "v \ of_complex x" "poincare_between 0\<^sub>h (of_complex x) v" hence "v \ circline_set x_axis" using poincare_between_poincare_line_uvz[of "0\<^sub>h" "of_complex x" v] using poincare_line_0_real_is_x_axis[of "of_complex x"] using \of_complex x \ 0\<^sub>h\ \v \ 0\<^sub>h\ \v \ of_complex x\ \of_complex x \ unit_disc\ \of_complex x \ circline_set x_axis\ by auto obtain v' where "v = of_complex v'" using \v \ unit_disc\ using inf_or_of_complex[of v] by auto hence **: "v = of_complex v'" "-1 < Re v'" "Re v' < 1" "Re v' \ 0" "is_real v'" using \v \ unit_disc\ \v \ 0\<^sub>h\ \v \ circline_set x_axis\ of_complex_inj[of v'] unfolding circline_set_x_axis by (auto simp add: cmod_eq_Re real_imag_0) show "let u' = to_complex (of_complex x); v' = to_complex v in arg u' = arg v' \ cmod u' \ cmod v'" using poincare_between_x_axis_0uv[of "Re x" "Re v'"] * ** using \poincare_between 0\<^sub>h (of_complex x) v\ using arg_complex_of_real_positive[of "Re x"] arg_complex_of_real_negative[of "Re x"] using arg_complex_of_real_positive[of "Re v'"] arg_complex_of_real_negative[of "Re v'"] by (auto simp add: cmod_eq_Re) next fix v assume "v \ unit_disc" "v \ 0\<^sub>h" "v \ of_complex x" then obtain v' where **: "v = of_complex v'" "v' \ 0" "v' \ x" using inf_or_of_complex[of v] by auto blast assume "let u' = to_complex (of_complex x); v' = to_complex v in arg u' = arg v' \ cmod u' \ cmod v'" hence ***: "Re x < 0 \ Re v' < 0 \ Re v' \ Re x \ 0 < Re x \ 0 < Re v' \ Re x \ Re v'" "is_real v'" using arg_pi_iff[of x] arg_pi_iff[of v'] using arg_0_iff[of x] arg_0_iff[of v'] using * ** by (smt cmod_Re_le_iff to_complex_of_complex)+ have "-1 < Re v'" "Re v' < 1" "Re v' \ 0" "is_real v'" using \v \ unit_disc\ ** \is_real v'\ by (auto simp add: cmod_eq_Re complex_eq_if_Re_eq) thus "poincare_between 0\<^sub>h (of_complex x) v" using poincare_between_x_axis_0uv[of "Re x" "Re v'"] * ** *** by simp qed next fix \ u assume "u \ unit_disc" "u \ 0\<^sub>h" let ?M = "moebius_rotation \" assume *: "?P' (moebius_pt ?M u)" show "?P' u" proof (rule allI, rule impI, (erule conjE)+) fix v assume "v \ unit_disc" "v \ 0\<^sub>h" "v \ u" have "moebius_pt ?M v \ moebius_pt ?M u" using \v \ u\ by auto obtain u' v' where "v = of_complex v'" "u = of_complex u'" "v' \ 0" "u' \ 0" using inf_or_of_complex[of u] inf_or_of_complex[of v] using \v \ unit_disc\ \u \ unit_disc\ \v \ 0\<^sub>h\ \u \ 0\<^sub>h\ by auto thus "?P u v" using *[rule_format, of "moebius_pt ?M v"] using \moebius_pt ?M v \ moebius_pt ?M u\ using unit_disc_fix_moebius_preserve_poincare_between[of ?M "0\<^sub>h" u v] using \v \ unit_disc\ \u \ unit_disc\ \v \ 0\<^sub>h\ \u \ 0\<^sub>h\ using arg_mult_eq[of "cis \" u' v'] by simp (auto simp add: arg_mult norm_mult) qed qed thus ?thesis using assms False by auto qed lemma poincare_between_y_axis_0uv: fixes x y :: real assumes "-1 < x" and "x < 1" and "x \ 0" assumes "-1 < y" and "y < 1" and "y \ 0" shows "poincare_between 0\<^sub>h (of_complex (\ * x)) (of_complex (\ * y)) \ (x < 0 \ y < 0 \ y \ x) \ (x > 0 \ y > 0 \ x \ y)" (is "?lhs \ ?rhs") using assms using poincare_between_0uv[of "of_complex (\ * x)" "of_complex (\ * y)"] using arg_pi2_iff[of "\ * cor x"] arg_pi2_iff[of "\ * cor y"] using arg_minus_pi2_iff[of "\ * cor x"] arg_minus_pi2_iff[of "\ * cor y"] apply (simp add: norm_mult) apply (smt (verit, best)) done lemma poincare_between_x_axis_uvw: fixes x y z :: real assumes "-1 < x" and "x < 1" assumes "-1 < y" and "y < 1" and "y \ x" assumes "-1 < z" and "z < 1" and "z \ x" shows "poincare_between (of_complex x) (of_complex y) (of_complex z) \ (y < x \ z < x \ z \ y) \ (y > x \ z > x \ y \ z)" (is "?lhs \ ?rhs") proof (cases "x = 0 \ y = 0 \ z = 0") case True thus ?thesis proof (cases "x = 0") case True thus ?thesis using poincare_between_x_axis_0uv assms by simp next case False show ?thesis proof (cases "z = 0") case True thus ?thesis using poincare_between_x_axis_0uv assms poincare_between_rev by (smt norm_of_real of_complex_zero of_real_0 poincare_between_nonstrict(2) unit_disc_iff_cmod_lt_1) next case False have "y = 0" using \x \ 0\ \z \ 0\ \x = 0 \ y = 0 \ z = 0\ by simp have "poincare_between (of_complex x) 0\<^sub>h (of_complex z) = (is_real z \ x * z < 0)" using \x \ 0\ \z \ 0\ poincare_between_x_axis_u0v by auto moreover have "x * z < 0 \ ?rhs" using True \x \ 0\ \z \ 0\ by (smt zero_le_mult_iff) ultimately show ?thesis using \y = 0\ by auto qed qed next case False thus ?thesis proof (cases "z = y") case True thus ?thesis using assms unfolding poincare_between_def by auto next case False let ?x = "of_complex x" and ?y = "of_complex y" and ?z = "of_complex z" have "?x \ unit_disc" "?y \ unit_disc" "?z \ unit_disc" using assms by auto have distinct: "distinct [?x, ?y, ?z, inversion ?y]" using \y \ x\ \z \ x\ False \?x \ unit_disc\ \?y \ unit_disc\ \?z \ unit_disc\ using inversion_noteq_unit_disc[of ?y ?y] using inversion_noteq_unit_disc[of ?y ?x] using inversion_noteq_unit_disc[of ?y ?z] using of_complex_inj[of x y] of_complex_inj[of y z] of_complex_inj[of x z] by auto have "cor y * cor x \ 1" using assms by (smt minus_mult_minus mult_less_cancel_left2 mult_less_cancel_right2 of_real_1 of_real_eq_iff of_real_mult) let ?cr = "cross_ratio ?x ?y ?z (inversion ?y)" have "Re (to_complex ?cr) = (x - y) * (z*y - 1)/ ((x*y - 1)*(z - y))" proof- have " \y x z. \y \ x; z \ x; z \ y; cor y * cor x \ 1; x \ 0; y \ 0; z \ 0\ \ (y * y + y * (y * (x * z)) - (y * x + y * (y * (y * z)))) / (y * y + y * (y * (x * z)) - (y * z + y * (y * (y * x)))) = (y + y * (x * z) - (x + y * (y * z))) / (y + y * (x * z) - (z + y * (y * x)))" by (metis (no_types, hide_lams) ab_group_add_class.ab_diff_conv_add_uminus distrib_left mult_divide_mult_cancel_left_if mult_minus_right) thus ?thesis using \y \ x\ \z \ x\ False \\ (x = 0 \ y = 0 \ z = 0)\ using \cor y * cor x \ 1\ unfolding inversion_def by (transfer, transfer, auto simp add: vec_cnj_def power2_eq_square field_simps split: if_split_asm) qed moreover have "(x*y - 1) < 0" using assms by (smt minus_mult_minus mult_less_cancel_right2 zero_less_mult_iff) moreover have "(z*y - 1) < 0" using assms by (smt minus_mult_minus mult_less_cancel_right2 zero_less_mult_iff) moreover have "(x - y) / (z - y) < 0 \ ?rhs" using \y \ x\ \z \ x\ False \\ (x = 0 \ y = 0 \ z = 0)\ by (smt divide_less_cancel divide_nonneg_nonpos divide_nonneg_pos divide_nonpos_nonneg divide_nonpos_nonpos) ultimately have *: "Re (to_complex ?cr) < 0 \ ?rhs" by (smt algebra_split_simps(24) minus_divide_left zero_less_divide_iff zero_less_mult_iff) show ?thesis proof assume ?lhs have "is_real (to_complex ?cr)" "Re (to_complex ?cr) < 0" using \?lhs\ distinct unfolding poincare_between_def Let_def by auto thus ?rhs using * by simp next assume ?rhs hence "Re (to_complex ?cr) < 0" using * by simp moreover have "{of_complex (cor x), of_complex (cor y), of_complex (cor z), inversion (of_complex (cor y))} \ circline_set x_axis" using \\ (x = 0 \ y = 0 \ z = 0)\ is_real_inversion[of "cor y"] using inf_or_of_complex[of "inversion ?y"] by (auto simp del: inversion_of_complex) hence "is_real (to_complex ?cr)" using four_points_on_circline_iff_cross_ratio_real[OF distinct] by auto ultimately show ?lhs using distinct unfolding poincare_between_def Let_def by auto qed qed qed subsubsection \H-betweenness and h-collinearity\ text\For three h-collinear points at least one of the three possible h-betweeness relations must hold.\ lemma poincare_collinear3_between: assumes "u \ unit_disc" and "v \ unit_disc" and "w \ unit_disc" assumes "poincare_collinear {u, v, w}" shows "poincare_between u v w \ poincare_between u w v \ poincare_between v u w" (is "?P' u v w") proof (cases "u=v") case True thus ?thesis using assms by auto next case False have "\ w. w \ unit_disc \ poincare_collinear {u, v, w} \ ?P' u v w" (is "?P u v") proof (rule wlog_positive_x_axis[where P="?P"]) fix x assume x: "is_real x" "0 < Re x" "Re x < 1" hence "x \ 0" using complex.expand[of x 0] by auto hence *: "poincare_line 0\<^sub>h (of_complex x) = x_axis" using x poincare_line_0_real_is_x_axis[of "of_complex x"] unfolding circline_set_x_axis by auto have "of_complex x \ unit_disc" using x by (auto simp add: cmod_eq_Re) have "of_complex x \ 0\<^sub>h" using \x \ 0\ by auto show "?P 0\<^sub>h (of_complex x)" proof safe fix w assume "w \ unit_disc" assume "poincare_collinear {0\<^sub>h, of_complex x, w}" hence "w \ circline_set x_axis" using * unique_poincare_line[of "0\<^sub>h" "of_complex x"] \of_complex x \ unit_disc\ \x \ 0\ \of_complex x \ 0\<^sub>h\ unfolding poincare_collinear_def by auto then obtain w' where w': "w = of_complex w'" "is_real w'" using \w \ unit_disc\ using inf_or_of_complex[of w] unfolding circline_set_x_axis by auto hence "-1 < Re w'" "Re w' < 1" using \w \ unit_disc\ by (auto simp add: cmod_eq_Re) assume 1: "\ poincare_between (of_complex x) 0\<^sub>h w" hence "w \ 0\<^sub>h" "w' \ 0" using w' unfolding poincare_between_def by auto hence "Re w' \ 0" using w' complex.expand[of w' 0] by auto have "Re w' \ 0" using 1 poincare_between_x_axis_u0v[of x w'] \Re x > 0\ \is_real x\ \x \ 0\ \w' \ 0\ w' using mult_pos_neg by force moreover assume "\ poincare_between 0\<^sub>h (of_complex x) w" hence "Re w' < Re x" using poincare_between_x_axis_0uv[of "Re x" "Re w'"] using w' x \-1 < Re w'\ \Re w' < 1\ \Re w' \ 0\ by auto ultimately show "poincare_between 0\<^sub>h w (of_complex x)" using poincare_between_x_axis_0uv[of "Re w'" "Re x"] using w' x \-1 < Re w'\ \Re w' < 1\ \Re w' \ 0\ by auto qed next show "u \ unit_disc" "v \ unit_disc" "u \ v" by fact+ next fix M u v assume 1: "unit_disc_fix M" "u \ unit_disc" "v \ unit_disc" "u \ v" let ?Mu = "moebius_pt M u" and ?Mv = "moebius_pt M v" assume 2: "?P ?Mu ?Mv" show "?P u v" proof safe fix w assume "w \ unit_disc" "poincare_collinear {u, v, w}" "\ poincare_between u v w" "\ poincare_between v u w" thus "poincare_between u w v" using 1 2[rule_format, of "moebius_pt M w"] by simp qed qed thus ?thesis using assms by simp qed lemma poincare_collinear3_iff: assumes "u \ unit_disc" "v \ unit_disc" "w \ unit_disc" shows "poincare_collinear {u, v, w} \ poincare_between u v w \ poincare_between v u w \ poincare_between v w u" using assms by (metis poincare_collinear3_between insert_commute poincare_between_poincare_collinear poincare_between_rev) subsection \Some properties of betweenness\ lemma poincare_between_transitivity: assumes "a \ unit_disc" and "x \ unit_disc" and "b \ unit_disc" and "y \ unit_disc" and "poincare_between a x b" and "poincare_between a b y" shows "poincare_between x b y" proof(cases "a = b") case True thus ?thesis using assms using poincare_between_sandwich by blast next case False have "\ x. \ y. poincare_between a x b \ poincare_between a b y \ x \ unit_disc \ y \ unit_disc \ poincare_between x b y" (is "?P a b") proof (rule wlog_positive_x_axis[where P="?P"]) show "a \ unit_disc" using assms by simp next show "b \ unit_disc" using assms by simp next show "a \ b" using False by simp next fix M u v assume *: "unit_disc_fix M" "u \ unit_disc" "v \ unit_disc" "u \ v" "\x y. poincare_between (moebius_pt M u) x (moebius_pt M v) \ poincare_between (moebius_pt M u) (moebius_pt M v) y \ x \ unit_disc \ y \ unit_disc \ poincare_between x (moebius_pt M v) y" show "\x y. poincare_between u x v \ poincare_between u v y \ x \ unit_disc \ y \ unit_disc \ poincare_between x v y" proof safe fix x y assume "poincare_between u x v" "poincare_between u v y" " x \ unit_disc" "y \ unit_disc" have "poincare_between (moebius_pt M u) (moebius_pt M x) (moebius_pt M v)" using \poincare_between u x v\ \unit_disc_fix M\ \x \ unit_disc\ \u \ unit_disc\ \v \ unit_disc\ by simp moreover have "poincare_between (moebius_pt M u) (moebius_pt M v) (moebius_pt M y)" using \poincare_between u v y\ \unit_disc_fix M\ \y \ unit_disc\ \u \ unit_disc\ \v \ unit_disc\ by simp moreover have "(moebius_pt M x) \ unit_disc" using \unit_disc_fix M\ \x \ unit_disc\ by simp moreover have "(moebius_pt M y) \ unit_disc" using \unit_disc_fix M\ \y \ unit_disc\ by simp ultimately have "poincare_between (moebius_pt M x) (moebius_pt M v) (moebius_pt M y)" using * by blast thus "poincare_between x v y" using \y \ unit_disc\ * \x \ unit_disc\ by simp qed next fix x assume xx: "is_real x" "0 < Re x" "Re x < 1" hence "of_complex x \ unit_disc" using cmod_eq_Re by auto hence "of_complex x \ \\<^sub>h" by simp have " of_complex x \ 0\<^sub>h" using xx by auto have "of_complex x \ circline_set x_axis" using xx by simp show "\m n. poincare_between 0\<^sub>h m (of_complex x) \ poincare_between 0\<^sub>h (of_complex x) n \ m \ unit_disc \ n \ unit_disc \ poincare_between m (of_complex x) n" proof safe fix m n assume **: "poincare_between 0\<^sub>h m (of_complex x)" "poincare_between 0\<^sub>h (of_complex x) n" "m \ unit_disc" " n \ unit_disc" show "poincare_between m (of_complex x) n" proof(cases "m = 0\<^sub>h") case True thus ?thesis using ** by auto next case False hence "m \ circline_set x_axis" using poincare_between_poincare_line_uzv[of "0\<^sub>h" "of_complex x" m] using poincare_line_0_real_is_x_axis[of "of_complex x"] using \of_complex x \ unit_disc\ \of_complex x \ \\<^sub>h\ \of_complex x \ 0\<^sub>h\ using \of_complex x \ circline_set x_axis\ \m \ unit_disc\ **(1) by simp then obtain m' where "m = of_complex m'" "is_real m'" using inf_or_of_complex[of m] \m \ unit_disc\ unfolding circline_set_x_axis by auto hence "Re m' \ Re x" using \poincare_between 0\<^sub>h m (of_complex x)\ xx \of_complex x \ 0\<^sub>h\ using False ** \of_complex x \ unit_disc\ using cmod_Re_le_iff poincare_between_0uv by auto have "n \ 0\<^sub>h" using **(2, 4) \of_complex x \ 0\<^sub>h\ \of_complex x \ unit_disc\ using poincare_between_sandwich by fastforce have "n \ circline_set x_axis" using poincare_between_poincare_line_uvz[of "0\<^sub>h" "of_complex x" n] using poincare_line_0_real_is_x_axis[of "of_complex x"] using \of_complex x \ unit_disc\ \of_complex x \ \\<^sub>h\ \of_complex x \ 0\<^sub>h\ using \of_complex x \ circline_set x_axis\ \n \ unit_disc\ **(2) by simp then obtain n' where "n = of_complex n'" "is_real n'" using inf_or_of_complex[of n] \n \ unit_disc\ unfolding circline_set_x_axis by auto hence "Re x \ Re n'" using \poincare_between 0\<^sub>h (of_complex x) n\ xx \of_complex x \ 0\<^sub>h\ using False ** \of_complex x \ unit_disc\ \n \ 0\<^sub>h\ using cmod_Re_le_iff poincare_between_0uv - by (metis Re_complex_of_real arg_0_iff rcis_cmod_arg rcis_zero_arg to_complex_of_complex) + by (metis Re_complex_of_real arg_0_iff rcis_cmod_Arg rcis_zero_arg to_complex_of_complex) have "poincare_between (of_complex m') (of_complex x) (of_complex n')" using \Re x \ Re n'\ \Re m' \ Re x\ using poincare_between_x_axis_uvw[of "Re m'" "Re x" "Re n'"] using \is_real n'\ \is_real m'\ \n \ unit_disc\ \n = of_complex n'\ using xx \m = of_complex m'\ \m \ unit_disc\ by (smt complex_of_real_Re norm_of_real poincare_between_def unit_disc_iff_cmod_lt_1) thus ?thesis using \n = of_complex n'\ \m = of_complex m'\ by auto qed qed qed thus ?thesis using assms by blast qed (* ------------------------------------------------------------------ *) subsection\Poincare between - sum distances\ (* ------------------------------------------------------------------ *) text\Another possible definition of the h-betweenness relation is given in terms of h-distances between pairs of points. We prove it as a characterization equivalent to our cross-ratio based definition.\ lemma poincare_between_sum_distances_x_axis_u0v: assumes "of_complex u' \ unit_disc" "of_complex v' \ unit_disc" assumes "is_real u'" "u' \ 0" "v' \ 0" shows "poincare_distance (of_complex u') 0\<^sub>h + poincare_distance 0\<^sub>h (of_complex v') = poincare_distance (of_complex u') (of_complex v') \ is_real v' \ Re u' * Re v' < 0" (is "?P u' v' \ ?Q u' v'") proof- have "Re u' \ 0" using \is_real u'\ \u' \ 0\ using complex_eq_if_Re_eq by simp let ?u = "cmod u'" and ?v = "cmod v'" and ?uv = "cmod (u' - v')" have disc: "?u\<^sup>2 < 1" "?v\<^sup>2 < 1" using unit_disc_cmod_square_lt_1[OF assms(1)] using unit_disc_cmod_square_lt_1[OF assms(2)] by auto have "poincare_distance (of_complex u') 0\<^sub>h + poincare_distance 0\<^sub>h (of_complex v') = arcosh (((1 + ?u\<^sup>2) * (1 + ?v\<^sup>2) + 4 * ?u * ?v) / ((1 - ?u\<^sup>2) * (1 - ?v\<^sup>2)))" (is "_ = arcosh ?r1") using poincare_distance_formula_zero_sum[OF assms(1-2)] by (simp add: Let_def) moreover have "poincare_distance (of_complex u') (of_complex v') = arcosh (((1 - ?u\<^sup>2) * (1 - ?v\<^sup>2) + 2 * ?uv\<^sup>2) / ((1 - ?u\<^sup>2) * (1 - ?v\<^sup>2)))" (is "_ = arcosh ?r2") using disc using poincare_distance_formula[OF assms(1-2)] by (subst add_divide_distrib) simp moreover have "arcosh ?r1 = arcosh ?r2 \ ?Q u' v'" proof assume "arcosh ?r1 = arcosh ?r2" hence "?r1 = ?r2" proof (subst (asm) arcosh_eq_iff) show "?r1 \ 1" proof- have "(1 - ?u\<^sup>2) * (1 - ?v\<^sup>2) \ (1 + ?u\<^sup>2) * (1 + ?v\<^sup>2) + 4 * ?u * ?v" by (simp add: field_simps) thus ?thesis using disc by simp qed next show "?r2 \ 1" using disc by simp qed hence "(1 + ?u\<^sup>2) * (1 + ?v\<^sup>2) + 4 * ?u * ?v = (1 - ?u\<^sup>2) * (1 - ?v\<^sup>2) + 2 * ?uv\<^sup>2" using disc by auto hence "(cmod (u' - v'))\<^sup>2 = (cmod u' + cmod v')\<^sup>2" by (simp add: field_simps power2_eq_square) hence *: "Re u' * Re v' + \Re u'\ * sqrt ((Im v')\<^sup>2 + (Re v')\<^sup>2) = 0" using \is_real u'\ unfolding cmod_power2 cmod_def by (simp add: field_simps) (simp add: power2_eq_square field_simps) hence "sqrt ((Im v')\<^sup>2 + (Re v')\<^sup>2) = \Re v'\" using \Re u' \ 0\ \v' \ 0\ by (smt complex_neq_0 mult.commute mult_cancel_right mult_minus_left real_sqrt_gt_0_iff) hence "Im v' = 0" by (smt Im_eq_0 norm_complex_def) moreover hence "Re u' * Re v' = - \Re u'\ * \Re v'\" using * by simp hence "Re u' * Re v' < 0" using \Re u' \ 0\ \v' \ 0\ by (simp add: \is_real v'\ complex_eq_if_Re_eq) ultimately show "?Q u' v'" by simp next assume "?Q u' v'" hence "is_real v'" "Re u' * Re v' < 0" by auto have "?r1 = ?r2" proof (cases "Re u' > 0") case True hence "Re v' < 0" using \Re u' * Re v' < 0\ by (smt zero_le_mult_iff) show ?thesis using disc \is_real u'\ \is_real v'\ using \Re u' > 0\ \Re v' < 0\ unfolding cmod_power2 cmod_def by simp (simp add: power2_eq_square field_simps) next case False hence "Re u' < 0" using \Re u' \ 0\ by simp hence "Re v' > 0" using \Re u' * Re v' < 0\ by (smt zero_le_mult_iff) show ?thesis using disc \is_real u'\ \is_real v'\ using \Re u' < 0\ \Re v' > 0\ unfolding cmod_power2 cmod_def by simp (simp add: power2_eq_square field_simps) qed thus "arcosh ?r1 = arcosh ?r2" by metis qed ultimately show ?thesis by simp qed text\ Different proof of the previous theorem relying on the cross-ratio definition, and not the distance formula. We suppose that this could be also used to prove the triangle inequality. \ lemma poincare_between_sum_distances_x_axis_u0v_different_proof: assumes "of_complex u' \ unit_disc" "of_complex v' \ unit_disc" assumes "is_real u'" "u' \ 0" "v' \ 0" (* additional condition *) "is_real v'" shows "poincare_distance (of_complex u') 0\<^sub>h + poincare_distance 0\<^sub>h (of_complex v') = poincare_distance (of_complex u') (of_complex v') \ Re u' * Re v' < 0" (is "?P u' v' \ ?Q u' v'") proof- have "-1 < Re u'" "Re u' < 1" "Re u' \ 0" using assms by (auto simp add: cmod_eq_Re complex_eq_if_Re_eq) have "-1 < Re v'" "Re v' < 1" "Re v' \ 0" using assms by (auto simp add: cmod_eq_Re complex_eq_if_Re_eq) have "\ln (Re ((1 - u') / (1 + u')))\ + \ln (Re ((1 - v') / (1 + v')))\ = \ln (Re ((1 + u') * (1 - v') / ((1 - u') * (1 + v'))))\ \ Re u' * Re v' < 0" (is "\ln ?a1\ + \ln ?a2\ = \ln ?a3\ \ _") proof- have 1: "0 < ?a1" "ln ?a1 > 0 \ Re u' < 0" using \Re u' < 1\ \Re u' > -1\ \is_real u'\ using complex_is_Real_iff by auto have 2: "0 < ?a2" "ln ?a2 > 0 \ Re v' < 0" using \Re v' < 1\ \Re v' > -1\ \is_real v'\ using complex_is_Real_iff by auto have 3: "0 < ?a3" "ln ?a3 > 0 \ Re v' < Re u'" using \Re u' < 1\ \Re u' > -1\ \is_real u'\ using \Re v' < 1\ \Re v' > -1\ \is_real v'\ using complex_is_Real_iff by auto (simp add: field_simps)+ show ?thesis proof assume *: "Re u' * Re v' < 0" show "\ln ?a1\ + \ln ?a2\ = \ln ?a3\" proof (cases "Re u' > 0") case True hence "Re v' < 0" using * by (smt mult_nonneg_nonneg) show ?thesis using 1 2 3 \Re u' > 0\ \Re v' < 0\ using \Re u' < 1\ \Re u' > -1\ \is_real u'\ using \Re v' < 1\ \Re v' > -1\ \is_real v'\ using complex_is_Real_iff using ln_div ln_mult by simp next case False hence "Re v' > 0" "Re u' < 0" using * by (smt zero_le_mult_iff)+ show ?thesis using 1 2 3 \Re u' < 0\ \Re v' > 0\ using \Re u' < 1\ \Re u' > -1\ \is_real u'\ using \Re v' < 1\ \Re v' > -1\ \is_real v'\ using complex_is_Real_iff using ln_div ln_mult by simp qed next assume *: "\ln ?a1\ + \ln ?a2\ = \ln ?a3\" { assume "Re u' > 0" "Re v' > 0" hence False using * 1 2 3 using \Re u' < 1\ \Re u' > -1\ \is_real u'\ using \Re v' < 1\ \Re v' > -1\ \is_real v'\ using complex_is_Real_iff using ln_mult ln_div by (cases "Re v' < Re u'") auto } moreover { assume "Re u' < 0" "Re v' < 0" hence False using * 1 2 3 using \Re u' < 1\ \Re u' > -1\ \is_real u'\ using \Re v' < 1\ \Re v' > -1\ \is_real v'\ using complex_is_Real_iff using ln_mult ln_div by (cases "Re v' < Re u'") auto } ultimately show "Re u' * Re v' < 0" using \Re u' \ 0\ \Re v' \ 0\ by (smt divisors_zero mult_le_0_iff) qed qed thus ?thesis using assms apply (subst poincare_distance_sym, simp, simp) apply (subst poincare_distance_zero_x_axis, simp, simp add: circline_set_x_axis) apply (subst poincare_distance_zero_x_axis, simp, simp add: circline_set_x_axis) apply (subst poincare_distance_x_axis_x_axis, simp, simp, simp add: circline_set_x_axis, simp add: circline_set_x_axis) apply simp done qed lemma poincare_between_sum_distances: assumes "u \ unit_disc" and "v \ unit_disc" and "w \ unit_disc" shows "poincare_between u v w \ poincare_distance u v + poincare_distance v w = poincare_distance u w" (is "?P' u v w") proof (cases "u = v") case True thus ?thesis using assms by simp next case False have "\ w. w \ unit_disc \ (poincare_between u v w \ poincare_distance u v + poincare_distance v w = poincare_distance u w)" (is "?P u v") proof (rule wlog_positive_x_axis) fix x assume "is_real x" "0 < Re x" "Re x < 1" have "of_complex x \ circline_set x_axis" using \is_real x\ by (auto simp add: circline_set_x_axis) have "of_complex x \ unit_disc" using \is_real x\ \0 < Re x\ \Re x < 1\ by (simp add: cmod_eq_Re) have "x \ 0" using \is_real x\ \Re x > 0\ by auto show "?P (of_complex x) 0\<^sub>h" proof (rule allI, rule impI) fix w assume "w \ unit_disc" then obtain w' where "w = of_complex w'" using inf_or_of_complex[of w] by auto show "?P' (of_complex x) 0\<^sub>h w" proof (cases "w = 0\<^sub>h") case True thus ?thesis by simp next case False hence "w' \ 0" using \w = of_complex w'\ by auto show ?thesis using \is_real x\ \x \ 0\ \w = of_complex w'\ \w' \ 0\ using \of_complex x \ unit_disc\ \w \ unit_disc\ apply simp apply (subst poincare_between_x_axis_u0v, simp_all) apply (subst poincare_between_sum_distances_x_axis_u0v, simp_all) done qed qed next show "v \ unit_disc" "u \ unit_disc" using assms by auto next show "v \ u" using \u \ v\ by simp next fix M u v assume *: "unit_disc_fix M" "u \ unit_disc" "v \ unit_disc" "u \ v" and **: "?P (moebius_pt M v) (moebius_pt M u)" show "?P v u" proof (rule allI, rule impI) fix w assume "w \ unit_disc" hence "moebius_pt M w \ unit_disc" using \unit_disc_fix M\ by auto thus "?P' v u w" using \u \ unit_disc\ \v \ unit_disc\ \w \ unit_disc\ \unit_disc_fix M\ using **[rule_format, of "moebius_pt M w"] by auto qed qed thus ?thesis using assms by simp qed subsection \Some more properties of h-betweenness.\ text \Some lemmas proved earlier are proved almost directly using the sum of distances characterization.\ lemma unit_disc_fix_moebius_preserve_poincare_between': assumes "unit_disc_fix M" and "u \ unit_disc" and "v \ unit_disc" and "w \ unit_disc" shows "poincare_between (moebius_pt M u) (moebius_pt M v) (moebius_pt M w) \ poincare_between u v w" using assms using poincare_between_sum_distances by simp lemma conjugate_preserve_poincare_between': assumes "u \ unit_disc" "v \ unit_disc" "w \ unit_disc" shows "poincare_between (conjugate u) (conjugate v) (conjugate w) \ poincare_between u v w" using assms using poincare_between_sum_distances by simp text \There is a unique point on a ray on the given distance from the given starting point\ lemma unique_poincare_distance_on_ray: assumes "d \ 0" "u \ v" "u \ unit_disc" "v \ unit_disc" assumes "y \ unit_disc" "poincare_distance u y = d" "poincare_between u v y" assumes "z \ unit_disc" "poincare_distance u z = d" "poincare_between u v z" shows "y = z" proof- have "\ d y z. d \ 0 \ y \ unit_disc \ poincare_distance u y = d \ poincare_between u v y \ z \ unit_disc \ poincare_distance u z = d \ poincare_between u v z \ y = z" (is "?P u v") proof (rule wlog_positive_x_axis[where P="?P"]) fix x assume x: "is_real x" "0 < Re x" "Re x < 1" hence "x \ 0" using complex.expand[of x 0] by auto hence *: "poincare_line 0\<^sub>h (of_complex x) = x_axis" using x poincare_line_0_real_is_x_axis[of "of_complex x"] unfolding circline_set_x_axis by auto have "of_complex x \ unit_disc" using x by (auto simp add: cmod_eq_Re) have "arg x = 0" using x using arg_0_iff by blast show "?P 0\<^sub>h (of_complex x)" proof safe fix y z assume "y \ unit_disc" "z \ unit_disc" then obtain y' z' where yz: "y = of_complex y'" "z = of_complex z'" using inf_or_of_complex[of y] inf_or_of_complex[of z] by auto assume betw: "poincare_between 0\<^sub>h (of_complex x) y" "poincare_between 0\<^sub>h (of_complex x) z" hence "y \ 0\<^sub>h" "z \ 0\<^sub>h" using \x \ 0\ \of_complex x \ unit_disc\ \y \ unit_disc\ using poincare_between_sandwich[of "0\<^sub>h" "of_complex x"] using of_complex_zero_iff[of x] by force+ hence "arg y' = 0" "cmod y' \ cmod x" "arg z' = 0" "cmod z' \ cmod x" using poincare_between_0uv[of "of_complex x" y] poincare_between_0uv[of "of_complex x" z] using \of_complex x \ unit_disc\ \x \ 0\ \arg x = 0\ \y \ unit_disc\ \z \ unit_disc\ betw yz by (simp_all add: Let_def) hence *: "is_real y'" "is_real z'" "Re y' > 0" "Re z' > 0" using arg_0_iff[of y'] arg_0_iff[of z'] x \y \ 0\<^sub>h\ \z \ 0\<^sub>h\ yz by auto assume "poincare_distance 0\<^sub>h z = poincare_distance 0\<^sub>h y" "0 \ poincare_distance 0\<^sub>h y" thus "y = z" using * yz \y \ unit_disc\ \z \ unit_disc\ using unique_x_axis_poincare_distance_positive[of "poincare_distance 0\<^sub>h y"] by (auto simp add: cmod_eq_Re unit_disc_to_complex_inj) qed next show "u \ unit_disc" "v \ unit_disc" "u \ v" by fact+ next fix M u v assume *: "unit_disc_fix M" "u \ unit_disc" "v \ unit_disc" "u \ v" assume **: "?P (moebius_pt M u) (moebius_pt M v)" show "?P u v" proof safe fix d y z assume ***: "0 \ poincare_distance u y" "y \ unit_disc" "poincare_between u v y" "z \ unit_disc" "poincare_between u v z" "poincare_distance u z = poincare_distance u y" let ?Mu = "moebius_pt M u" and ?Mv = "moebius_pt M v" and ?My = "moebius_pt M y" and ?Mz = "moebius_pt M z" have "?Mu \ unit_disc" "?Mv \ unit_disc" "?My \ unit_disc" "?Mz \ unit_disc" using \u \ unit_disc\ \v \ unit_disc\ \y \ unit_disc\ \z \ unit_disc\ using \unit_disc_fix M\ by auto hence "?My = ?Mz" using * *** using **[rule_format, of "poincare_distance ?Mu ?My" ?My ?Mz] by simp thus "y = z" using bij_moebius_pt[of M] unfolding bij_def inj_on_def by blast qed qed thus ?thesis using assms by auto qed end \ No newline at end of file diff --git a/thys/Poincare_Disc/Poincare_Tarski.thy b/thys/Poincare_Disc/Poincare_Tarski.thy --- a/thys/Poincare_Disc/Poincare_Tarski.thy +++ b/thys/Poincare_Disc/Poincare_Tarski.thy @@ -1,3044 +1,3044 @@ section \Poincar\'e model satisfies Tarski axioms\ theory Poincare_Tarski imports Poincare Poincare_Lines_Axis_Intersections Tarski begin (* ------------------------------------------------------------------ *) subsection\Pasch axiom\ (* ------------------------------------------------------------------ *) lemma Pasch_fun_mono: fixes r1 r2 :: real assumes "0 < r1" and "r1 \ r2" and "r2 < 1" shows "r1 + 1/r1 \ r2 + 1/r2" proof (cases "r1 = r2") case True thus ?thesis by simp next case False hence "r2 - r1 > 0" using assms by simp have "r1 * r2 < 1" using assms by (smt mult_le_cancel_left1) hence "1 / (r1 * r2) > 1" using assms by simp hence "(r2 - r1) / (r1 * r2) > (r2 - r1)" using \r2 - r1 > 0\ using mult_less_cancel_left_pos[of "r2 - r1" 1 "1 / (r1 * r2)"] by simp hence "1 / r1 - 1 / r2 > r2 - r1" using assms by (simp add: field_simps) thus ?thesis by simp qed text\Pasch axiom, non-degenerative case.\ lemma Pasch_nondeg: assumes "x \ unit_disc" and "y \ unit_disc" and "z \ unit_disc" and "u \ unit_disc" and "v \ unit_disc" assumes "distinct [x, y, z, u, v]" assumes "\ poincare_collinear {x, y, z}" assumes "poincare_between x u z" and "poincare_between y v z" shows "\ a. a \ unit_disc \ poincare_between u a y \ poincare_between x a v" proof- have "\ y z u. distinct [x, y, z, u, v] \ \ poincare_collinear {x, y, z} \ y \ unit_disc \ z \ unit_disc \ u \ unit_disc \ poincare_between x u z \ poincare_between y v z \ (\ a. a \ unit_disc \ poincare_between u a y \ poincare_between x a v)" (is "?P x v") proof (rule wlog_positive_x_axis[where P="?P"]) fix v assume v: "is_real v" "0 < Re v" "Re v < 1" hence "of_complex v \ unit_disc" by (auto simp add: cmod_eq_Re) show "?P 0\<^sub>h (of_complex v)" proof safe fix y z u assume distinct: "distinct [0\<^sub>h, y, z, u, of_complex v]" assume in_disc: "y \ unit_disc" "z \ unit_disc" "u \ unit_disc" then obtain y' z' u' where *: "y = of_complex y'" "z = of_complex z'" "u = of_complex u'" using inf_or_of_complex inf_notin_unit_disc by metis have "y' \ 0" "z' \ 0" "u' \ 0" "v \ 0" "y' \ z'" "y' \ u'" "z' \ u'" "y \ z" "y \ u" "z \ u" using of_complex_inj distinct * by auto note distinct = distinct this assume "\ poincare_collinear {0\<^sub>h, y, z}" hence nondeg_yz: "y'*cnj z' \ cnj y' * z'" using * poincare_collinear_zero_iff[of y' z'] in_disc distinct by auto assume "poincare_between 0\<^sub>h u z" hence "arg u' = arg z'" "cmod u' \ cmod z'" using * poincare_between_0uv[of u z] distinct in_disc by auto then obtain \ ru rz where uz_polar: "u' = cor ru * cis \" "z' = cor rz * cis \" "0 < ru" "ru \ rz" "0 < rz" and "\ = arg u'" "\ = arg z'" using * \u' \ 0\ \z' \ 0\ by (smt cmod_cis norm_le_zero_iff) obtain \ ry where y_polar: "y' = cor ry * cis \" "ry > 0" and "\ = arg y'" using \y' \ 0\ by (smt cmod_cis norm_le_zero_iff) from in_disc * \u' = cor ru * cis \\ \z' = cor rz * cis \\ \y' = cor ry * cis \\ have "ru < 1" "rz < 1" "ry < 1" by (auto simp: norm_mult) note polar = this y_polar uz_polar have nondeg: "cis \ * cis (- \) \ cis (- \) * cis \" using nondeg_yz polar by simp let ?yz = "poincare_line y z" let ?v = "calc_x_axis_intersection ?yz" assume "poincare_between y (of_complex v) z" hence "of_complex v \ circline_set ?yz" using in_disc \of_complex v \ unit_disc\ using distinct poincare_between_poincare_collinear[of y "of_complex v" z] using unique_poincare_line[of y z] by (auto simp add: poincare_collinear_def) moreover have "of_complex v \ circline_set x_axis" using \is_real v\ unfolding circline_set_x_axis by auto moreover have "?yz \ x_axis" proof (rule ccontr) assume "\ ?thesis" hence "{0\<^sub>h, y, z} \ circline_set (poincare_line y z)" unfolding circline_set_def using distinct poincare_line[of y z] by auto hence "poincare_collinear {0\<^sub>h, y, z}" unfolding poincare_collinear_def using distinct by force thus False using \\ poincare_collinear {0\<^sub>h, y, z}\ by simp qed ultimately have "?v = of_complex v" "intersects_x_axis ?yz" using unique_calc_x_axis_intersection[of "poincare_line y z" "of_complex v"] using intersects_x_axis_iff[of ?yz] using distinct \of_complex v \ unit_disc\ by (metis IntI is_poincare_line_poincare_line)+ have "intersects_x_axis_positive ?yz" using \Re v > 0\ \of_complex v \ unit_disc\ using \of_complex v \ circline_set ?yz\ \of_complex v \ circline_set x_axis\ using intersects_x_axis_positive_iff[of ?yz] \y \ z\ \?yz \ x_axis\ unfolding positive_x_axis_def by force have "y \ circline_set x_axis" proof (rule ccontr) assume "\ ?thesis" moreover hence "poincare_line y (of_complex v) = x_axis" using distinct \of_complex v \ circline_set x_axis\ using in_disc \of_complex v \ unit_disc\ using unique_poincare_line[of y "of_complex v" x_axis] by simp moreover have "z \ circline_set (poincare_line y (of_complex v))" using \of_complex v \ circline_set ?yz\ using unique_poincare_line[of y "of_complex v" "poincare_line y z"] using in_disc \of_complex v \ unit_disc\ distinct using poincare_line[of y z] unfolding circline_set_def by (metis distinct_length_2_or_more is_poincare_line_poincare_line mem_Collect_eq) ultimately have "y \ circline_set x_axis" "z \ circline_set x_axis" by auto hence "poincare_collinear {0\<^sub>h, y, z}" unfolding poincare_collinear_def by force thus False using \\ poincare_collinear {0\<^sub>h, y, z}\ by simp qed moreover have "z \ circline_set x_axis" proof (rule ccontr) assume "\ ?thesis" moreover hence "poincare_line z (of_complex v) = x_axis" using distinct \of_complex v \ circline_set x_axis\ using in_disc \of_complex v \ unit_disc\ using unique_poincare_line[of z "of_complex v" x_axis] by simp moreover have "y \ circline_set (poincare_line z (of_complex v))" using \of_complex v \ circline_set ?yz\ using unique_poincare_line[of z "of_complex v" "poincare_line y z"] using in_disc \of_complex v \ unit_disc\ distinct using poincare_line[of y z] unfolding circline_set_def by (metis distinct_length_2_or_more is_poincare_line_poincare_line mem_Collect_eq) ultimately have "y \ circline_set x_axis" "z \ circline_set x_axis" by auto hence "poincare_collinear {0\<^sub>h, y, z}" unfolding poincare_collinear_def by force thus False using \\ poincare_collinear {0\<^sub>h, y, z}\ by simp qed ultimately have "\ * \ < 0" using \poincare_between y (of_complex v) z\ using poincare_between_x_axis_intersection[of y z "of_complex v"] using in_disc \of_complex v \ unit_disc\ distinct using \of_complex v \ circline_set ?yz\ \of_complex v \ circline_set x_axis\ using \\ = arg z'\ \\ = arg y'\ * by (simp add: field_simps) have "\ \ pi" "\ \ 0" using \z \ circline_set x_axis\ * polar cis_pi unfolding circline_set_x_axis by auto have "\ \ pi" "\ \ 0" using \y \ circline_set x_axis\ * polar cis_pi unfolding circline_set_x_axis by auto have phi_sin: "\ > 0 \ sin \ > 0" "\ < 0 \ sin \ < 0" using \\ = arg z'\ \\ \ 0\ \\ \ pi\ - using arg_bounded[of z'] + using Arg_bounded[of z'] by (smt sin_gt_zero sin_le_zero sin_pi_minus sin_0_iff_canon sin_ge_zero)+ have theta_sin: "\ > 0 \ sin \ > 0" "\ < 0 \ sin \ < 0" using \\ = arg y'\ \\ \ 0\ \\ \ pi\ - using arg_bounded[of y'] + using Arg_bounded[of y'] by (smt sin_gt_zero sin_le_zero sin_pi_minus sin_0_iff_canon sin_ge_zero)+ have "sin \ * sin \ < 0" using \\ * \ < 0\ phi_sin theta_sin by (simp add: mult_less_0_iff) have "sin (\ - \) \ 0" proof (rule ccontr) assume "\ ?thesis" hence "sin (\ - \) = 0" by simp have "- 2 * pi < \ - \" "\ - \ < 2 * pi" - using \\ = arg z'\ \\ = arg y'\ arg_bounded[of z'] arg_bounded[of y'] \\ \ pi\ \\ \ pi\ + using \\ = arg z'\ \\ = arg y'\ Arg_bounded[of z'] Arg_bounded[of y'] \\ \ pi\ \\ \ pi\ by auto hence "\ - \ = -pi \ \ - \ = 0 \ \ - \ = pi" using \sin (\ - \) = 0\ by (smt sin_0_iff_canon sin_periodic_pi2) moreover { assume "\ - \ = - pi" hence "\ = \ - pi" by simp hence False using nondeg_yz using \y' = cor ry * cis \\ \z' = cor rz * cis \\ \rz > 0\ \ry > 0\ by auto } moreover { assume "\ - \ = 0" hence "\ = \" by simp hence False using \y' = cor ry * cis \\ \z' = cor rz * cis \\ \rz > 0\ \ry > 0\ using nondeg_yz by auto } moreover { assume "\ - \ = pi" hence "\ = \ + pi" by simp hence False using \y' = cor ry * cis \\ \z' = cor rz * cis \\ \rz > 0\ \ry > 0\ using nondeg_yz by auto } ultimately show False by auto qed have "u \ circline_set x_axis" proof- have "\ is_real u'" using * polar in_disc using \\ \ 0\ \\ = arg u'\ \\ \ pi\ phi_sin(1) phi_sin(2) by (metis is_real_arg2) moreover have "u \ \\<^sub>h" using in_disc by auto ultimately show ?thesis using * of_complex_inj[of u'] unfolding circline_set_x_axis by auto qed let ?yu = "poincare_line y u" have nondeg_yu: "y' * cnj u' \ cnj u' * u'" using nondeg_yz polar \ru > 0\ \rz > 0\ distinct by auto { (* derive results simultaneously for both u and z *) fix r :: real assume "r > 0" have den: "cor ry * cis \ * cnj 1 * cnj (cor r * cis \) * 1 - cor r * cis \ * cnj 1 * cnj (cor ry * cis \) * 1 \ 0" using \0 < r\ \0 < ry\ nondeg by auto let ?A = "2 * r * ry * sin(\ - \)" let ?B = "\ * (r * cis \ * (1 + ry\<^sup>2) - ry * cis \ * (1 + r\<^sup>2))" let ?ReB = "ry * (1 + r\<^sup>2) * sin \ - r * (1 + ry\<^sup>2) * sin \" have "Re (\ * (r * cis (-\) * ry * cis (\) - ry * cis (-\) * r * cis (\))) = ?A" by (simp add: sin_diff field_simps) moreover have "cor ry * cis (- \) * (cor ry * cis \) = ry\<^sup>2" "cor r * cis (- \) * (cor r * cis \) = r\<^sup>2" by (metis cis_inverse cis_neq_zero divide_complex_def cor_squared nonzero_mult_div_cancel_right power2_eq_square semiring_normalization_rules(15))+ ultimately have 1: "poincare_line_cvec_cmat (of_complex_cvec (cor ry * cis \)) (of_complex_cvec (cor r * cis \)) = (?A, ?B, cnj ?B, ?A)" using den unfolding poincare_line_cvec_cmat_def of_complex_cvec_def Let_def prod.case by (simp add: field_simps) have 2: "is_real ?A" by simp let ?mix = "cis \ * cis (- \) - cis (- \) * cis \" have "is_imag ?mix" using eq_minus_cnj_iff_imag[of ?mix] by simp hence "Im ?mix \ 0" using nondeg using complex.expand[of ?mix 0] by auto hence 3: "Re ?A \ 0" using \r > 0\ \ry > 0\ by (simp add: sin_diff field_simps) have "?A \ 0" using 2 3 by auto hence 4: "cor ?A \ 0" using 2 3 by (metis zero_complex.simps(1)) have 5: "?ReB / ?A = (sin \) / (2 * sin(\ - \)) * (1/r + r) - (sin \) / (2 * sin (\ - \)) * (1/ry + ry)" using \ry > 0\ \r > 0\ apply (subst diff_divide_distrib) apply (subst add_frac_num, simp) apply (subst add_frac_num, simp) apply (simp add: power2_eq_square mult.commute) apply (simp add: field_simps) done have "poincare_line_cvec_cmat (of_complex_cvec (cor ry * cis \)) (of_complex_cvec (cor r * cis \)) = (?A, ?B, cnj ?B, ?A) \ is_real ?A \ Re ?A \ 0 \ ?A \ 0 \ cor ?A \ 0 \ Re ?B = ?ReB \ ?ReB / ?A = (sin \) / (2 * sin(\ - \)) * (1/r + r) - (sin \) / (2 * sin (\ - \)) * (1/ry + ry)" using 1 2 3 4 5 by auto } note ** = this let ?Ayz = "2 * rz * ry * sin (\ - \)" let ?Byz = "\ * (rz * cis \ * (1 + ry\<^sup>2) - ry * cis \ * (1 + rz\<^sup>2))" let ?ReByz = "ry * (1 + rz\<^sup>2) * sin \ - rz * (1 + ry\<^sup>2) * sin \" let ?Kz = "(sin \) / (2 * sin(\ - \)) * (1/rz + rz) - (sin \) / (2 * sin (\ - \)) * (1/ry + ry)" have yz: "poincare_line_cvec_cmat (of_complex_cvec (cor ry * cis \)) (of_complex_cvec (cor rz * cis \)) = (?Ayz, ?Byz, cnj ?Byz, ?Ayz)" "is_real ?Ayz" "Re ?Ayz \ 0" "?Ayz \ 0" "cor ?Ayz \ 0" "Re ?Byz = ?ReByz" and Kz: "?ReByz / ?Ayz = ?Kz" using **[OF \0 < rz\] by auto let ?Ayu = "2 * ru * ry * sin (\ - \)" let ?Byu = "\ * (ru * cis \ * (1 + ry\<^sup>2) - ry * cis \ * (1 + ru\<^sup>2))" let ?ReByu = "ry * (1 + ru\<^sup>2) * sin \ - ru * (1 + ry\<^sup>2) * sin \" let ?Ku = "(sin \) / (2 * sin(\ - \)) * (1/ru + ru) - (sin \) / (2 * sin (\ - \)) * (1/ry + ry)" have yu: "poincare_line_cvec_cmat (of_complex_cvec (cor ry * cis \)) (of_complex_cvec (cor ru * cis \)) = (?Ayu, ?Byu, cnj ?Byu, ?Ayu)" "is_real ?Ayu" "Re ?Ayu \ 0" "?Ayu \ 0" "cor ?Ayu \ 0" "Re ?Byu = ?ReByu" and Ku: "?ReByu / ?Ayu = ?Ku" using **[OF \0 < ru\] by auto have "?Ayz \ 0" using \sin (\ - \) \ 0\ \ry > 0\ \rz > 0\ by auto have "Re ?Byz / ?Ayz < -1" using \intersects_x_axis_positive ?yz\ * \y' = cor ry * cis \\ \z' = cor rz * cis \\ \u' = cor ru * cis \\ apply simp apply (transfer fixing: ry rz ru \ \) apply (transfer fixing: ry rz ru \ \) proof- assume "intersects_x_axis_positive_cmat (poincare_line_cvec_cmat (of_complex_cvec (cor ry * cis \)) (of_complex_cvec (cor rz * cis \)))" thus "(ry * sin \ * (1 + rz\<^sup>2) - rz * sin \ * (1 + ry\<^sup>2)) / (2 * rz * ry * sin (\ - \)) < - 1" using yz by simp qed have "?ReByz / ?Ayz \ ?ReByu / ?Ayu" proof (cases "sin \ > 0") case True hence "sin \ < 0" using \sin \ * sin \ < 0\ by (smt mult_nonneg_nonneg) have "?ReByz < 0" proof- have "ry * (1 + rz\<^sup>2) * sin \ < 0" using \ry > 0\ \rz > 0\ using \sin \ < 0\ by (smt mult_pos_neg mult_pos_pos zero_less_power) moreover have "rz * (1 + ry\<^sup>2) * sin \ > 0" using \ry > 0\ \rz > 0\ using \sin \ > 0\ by (smt mult_pos_neg mult_pos_pos zero_less_power) ultimately show ?thesis by simp qed have "?Ayz > 0" using \Re ?Byz / ?Ayz < -1\ \Re ?Byz = ?ReByz\ \?ReByz < 0\ by (smt divide_less_0_iff) hence "sin (\ - \) > 0" using \ry > 0\ \rz > 0\ by (smt mult_pos_pos zero_less_mult_pos) have "1 / ru + ru \ 1 / rz + rz" using Pasch_fun_mono[of ru rz] \0 < ru\ \ru \ rz\ \rz < 1\ by simp hence "sin \ * (1 / ru + ru) \ sin \ * (1 / rz + rz)" using \sin \ < 0\ by auto thus ?thesis using \ru > 0\ \rz > 0\ \ru \ rz\ \rz < 1\ \?Ayz > 0\ \sin (\ - \) > 0\ using divide_right_mono[of "sin \ * (1 / ru + ru)" "sin \ * (1 / rz + rz)" "2 * sin (\ - \)"] by (subst Kz, subst Ku) simp next assume "\ sin \ > 0" hence "sin \ < 0" using \sin \ * sin \ < 0\ by (cases "sin \ = 0", simp_all) hence "sin \ > 0" using \sin \ * sin \ < 0\ by (smt mult_nonpos_nonpos) have "?ReByz > 0" proof- have "ry * (1 + rz\<^sup>2) * sin \ > 0" using \ry > 0\ \rz > 0\ using \sin \ > 0\ by (smt mult_pos_neg mult_pos_pos zero_less_power) moreover have "rz * (1 + ry\<^sup>2) * sin \ < 0" using \ry > 0\ \rz > 0\ using \sin \ < 0\ by (smt mult_pos_neg mult_pos_pos zero_less_power) ultimately show ?thesis by simp qed have "?Ayz < 0" using \Re ?Byz / ?Ayz < -1\ \?Ayz \ 0\ \Re ?Byz = ?ReByz\ \?ReByz > 0\ by (smt divide_less_0_iff) hence "sin (\ - \) < 0" using \ry > 0\ \rz > 0\ by (smt mult_nonneg_nonneg) have "1 / ru + ru \ 1 / rz + rz" using Pasch_fun_mono[of ru rz] \0 < ru\ \ru \ rz\ \rz < 1\ by simp hence "sin \ * (1 / ru + ru) \ sin \ * (1 / rz + rz)" using \sin \ > 0\ by auto thus ?thesis using \ru > 0\ \rz > 0\ \ru \ rz\ \rz < 1\ \?Ayz < 0\ \sin (\ - \) < 0\ using divide_right_mono_neg[of "sin \ * (1 / rz + rz)" "sin \ * (1 / ru + ru)" "2 * sin (\ - \)"] by (subst Kz, subst Ku) simp qed have "intersects_x_axis_positive ?yu" using * \y' = cor ry * cis \\ \z' = cor rz * cis \\ \u' = cor ru * cis \\ apply simp apply (transfer fixing: ry rz ru \ \) apply (transfer fixing: ry rz ru \ \) proof- have "Re ?Byu / ?Ayu < -1" using \Re ?Byz / ?Ayz < -1\ \?ReByz / ?Ayz \ ?ReByu / ?Ayu\ by (subst (asm) \Re ?Byz = ?ReByz\, subst \Re ?Byu = ?ReByu\) simp thus "intersects_x_axis_positive_cmat (poincare_line_cvec_cmat (of_complex_cvec (cor ry * cis \)) (of_complex_cvec (cor ru * cis \)))" using yu by simp qed let ?a = "calc_x_axis_intersection ?yu" have "?a \ positive_x_axis" "?a \ circline_set ?yu" "?a \ unit_disc" using \intersects_x_axis_positive ?yu\ using intersects_x_axis_positive_iff'[of ?yu] \y \ u\ by auto then obtain a' where a': "?a = of_complex a'" "is_real a'" "Re a' > 0" "Re a' < 1" unfolding positive_x_axis_def circline_set_x_axis by (auto simp add: cmod_eq_Re) have "intersects_x_axis ?yz" "intersects_x_axis ?yu" using \intersects_x_axis_positive ?yz\ \intersects_x_axis_positive ?yu\ by auto show "\a. a \ unit_disc \ poincare_between u a y \ poincare_between 0\<^sub>h a (of_complex v)" proof (rule_tac x="?a" in exI, safe) show "poincare_between u ?a y" using poincare_between_x_axis_intersection[of y u ?a] using calc_x_axis_intersection[OF is_poincare_line_poincare_line[OF \y \ u\] \intersects_x_axis ?yu\] using calc_x_axis_intersection_in_unit_disc[OF is_poincare_line_poincare_line[OF \y \ u\] \intersects_x_axis ?yu\] using in_disc \y \ u\ \y \ circline_set x_axis\ \u \ circline_set x_axis\ using * \\ = arg u'\ \\ = arg y'\ \\ * \ < 0\ by (subst poincare_between_rev, auto simp add: mult.commute) next show "poincare_between 0\<^sub>h ?a (of_complex v)" proof- have "-?ReByz / ?Ayz \ -?ReByu / ?Ayu" using \?ReByz / ?Ayz \ ?ReByu / ?Ayu\ by linarith have "outward ?yz ?yu" using * \y' = cor ry * cis \\ \z' = cor rz * cis \\ \u' = cor ru * cis \\ apply simp apply (transfer fixing: ry rz ru \ \) apply (transfer fixing: ry rz ru \ \) apply (subst yz yu)+ unfolding outward_cmat_def apply (simp only: Let_def prod.case) apply (subst yz yu)+ using \-?ReByz / ?Ayz \ -?ReByu / ?Ayu\ by simp hence "Re a' \ Re v" using \?v = of_complex v\ using \?a = of_complex a'\ using \intersects_x_axis_positive ?yz\ \intersects_x_axis_positive ?yu\ using outward[OF is_poincare_line_poincare_line[OF \y \ z\] is_poincare_line_poincare_line[OF \y \ u\]] by simp thus ?thesis using \?v = of_complex v\ using poincare_between_x_axis_0uv[of "Re a'" "Re v"] a' v by simp qed next show "?a \ unit_disc" by fact qed qed next show "x \ unit_disc" "v \ unit_disc" "x \ v" using assms by auto next fix M x v let ?Mx = "moebius_pt M x" and ?Mv = "moebius_pt M v" assume 1: "unit_disc_fix M" "x \ unit_disc" "v \ unit_disc" "x \ v" assume 2: "?P ?Mx ?Mv" show "?P x v" proof safe fix y z u let ?My = "moebius_pt M y" and ?Mz = "moebius_pt M z" and ?Mu = "moebius_pt M u" assume "distinct [x, y, z, u, v]" "\ poincare_collinear {x, y, z}" "y \ unit_disc" "z \ unit_disc" "u \ unit_disc" "poincare_between x u z" "poincare_between y v z" hence "\ Ma. Ma \ unit_disc \ poincare_between ?Mu Ma ?My \ poincare_between ?Mx Ma ?Mv" using 1 2[rule_format, of ?My ?Mz ?Mu] by simp then obtain Ma where Ma: "Ma \ unit_disc" "poincare_between ?Mu Ma ?My \ poincare_between ?Mx Ma ?Mv" by blast let ?a = "moebius_pt (-M) Ma" let ?Ma = "moebius_pt M ?a" have "?Ma = Ma" by (metis moebius_pt_invert uminus_moebius_def) hence "?Ma \ unit_disc" "poincare_between ?Mu ?Ma ?My \ poincare_between ?Mx ?Ma ?Mv" using Ma by auto thus "\a. a \ unit_disc \ poincare_between u a y \ poincare_between x a v" using unit_disc_fix_moebius_inv[OF \unit_disc_fix M\] \unit_disc_fix M\ \Ma \ unit_disc\ using \u \ unit_disc\ \v \ unit_disc\ \x \ unit_disc\ \y \ unit_disc\ by (rule_tac x="?a" in exI, simp del: moebius_pt_comp_inv_right) qed qed thus ?thesis using assms by auto qed text\Pasch axiom, only degenerative cases.\ lemma Pasch_deg: assumes "x \ unit_disc" and "y \ unit_disc" and "z \ unit_disc" and "u \ unit_disc" and "v \ unit_disc" assumes "\ distinct [x, y, z, u, v] \ poincare_collinear {x, y, z}" assumes "poincare_between x u z" and "poincare_between y v z" shows "\ a. a \ unit_disc \ poincare_between u a y \ poincare_between x a v" proof(cases "poincare_collinear {x, y, z}") case True hence "poincare_between x y z \ poincare_between y x z \ poincare_between y z x" using assms(1, 2, 3) poincare_collinear3_between poincare_between_rev by blast show ?thesis proof(cases "poincare_between x y z") case True have "poincare_between x y v" using True assms poincare_between_transitivity by (meson poincare_between_rev) thus ?thesis using assms(2) by (rule_tac x="y" in exI, simp) next case False hence "poincare_between y x z \ poincare_between y z x" using \poincare_between x y z \ poincare_between y x z \ poincare_between y z x\ by simp show ?thesis proof(cases "poincare_between y x z") case True hence "poincare_between u x y" using assms by (meson poincare_between_rev poincare_between_transitivity) thus ?thesis using assms by (rule_tac x="x" in exI, simp) next case False hence "poincare_between y z x" using \poincare_between y x z \ poincare_between y z x\ by auto hence "poincare_between x z v" using assms by (meson poincare_between_rev poincare_between_transitivity) hence "poincare_between x u v" using assms poincare_between_transitivity poincare_between_rev by (smt poincare_between_sum_distances) thus ?thesis using assms by (rule_tac x="u" in exI, simp) qed qed next case False hence "\ distinct [x, y, z, u, v]" using assms(6) by auto show ?thesis proof(cases "u=z") case True thus ?thesis using assms apply(rule_tac x="v" in exI) by(simp add:poincare_between_rev) next case False (* "u \ z" *) hence "x \ z" using assms poincare_between_sandwich by blast show ?thesis proof(cases "v=z") case True thus ?thesis using assms by (rule_tac x="u" in exI, simp) next case False (* v \ z *) hence "y \ z" using assms poincare_between_sandwich by blast show ?thesis proof(cases "u = x") case True thus ?thesis using assms by (rule_tac x="x" in exI, simp) next case False (*u \ x*) have "x \ y" using assms \\ poincare_collinear {x, y, z}\ by fastforce have "x \ v" using assms \\ poincare_collinear {x, y, z}\ by (metis insert_commute poincare_between_poincare_collinear) have "u \ y" using assms \\ poincare_collinear {x, y, z}\ using poincare_between_poincare_collinear by blast have "u \ v" proof(rule ccontr) assume "\ u \ v" hence "poincare_between x v z" using assms by auto hence "x \ circline_set (poincare_line z v)" using poincare_between_rev[of x v z] using poincare_between_poincare_line_uvz[of z v x] using assms \v \ z\ by auto have "y \ circline_set (poincare_line z v)" using assms \\ u \ v\ using poincare_between_rev[of y v z] using poincare_between_poincare_line_uvz[of z v y] using assms \v \ z\ by auto have "z \ circline_set (poincare_line z v)" using ex_poincare_line_two_points[of z v] \v \ z\ by auto have "is_poincare_line (poincare_line z v)" using \v \ z\ by auto hence "poincare_collinear {x, y, z}" using \x \ circline_set (poincare_line z v)\ using \y \ circline_set (poincare_line z v)\ using \z \ circline_set (poincare_line z v)\ unfolding poincare_collinear_def by (rule_tac x="poincare_line z v" in exI, simp) thus False using \\ poincare_collinear {x, y, z}\ by simp qed have "v = y" using \u \ v\ \u \ y\ \x \ v\ \x \ y\ \u \ x\ \y \ z\ \v \ z\ \x \ z\ \u \ z\ using \\ distinct [x, y, z, u, v]\ by auto thus ?thesis using assms by (rule_tac x="y" in exI, simp) qed qed qed qed text \Axiom of Pasch\ lemma Pasch: assumes "x \ unit_disc" and "y \ unit_disc" and "z \ unit_disc" and "u \ unit_disc" and "v \ unit_disc" assumes "poincare_between x u z" and "poincare_between y v z" shows "\ a. a \ unit_disc \ poincare_between u a y \ poincare_between x a v" proof(cases "distinct [x, y, z, u, v] \ \ poincare_collinear {x, y, z}") case True thus ?thesis using assms Pasch_nondeg by auto next case False thus ?thesis using assms Pasch_deg by auto qed (* ------------------------------------------------------------------ *) subsection\Segment construction axiom\ (* ------------------------------------------------------------------ *) lemma segment_construction: assumes "x \ unit_disc" and "y \ unit_disc" assumes "a \ unit_disc" and "b \ unit_disc" shows "\ z. z \ unit_disc \ poincare_between x y z \ poincare_distance y z = poincare_distance a b" proof- obtain d where d: "d = poincare_distance a b" by auto have "d \ 0" using assms by (simp add: d poincare_distance_ge0) have "\ z. z \ unit_disc \ poincare_between x y z \ poincare_distance y z = d" (is "?P x y") proof (cases "x = y") case True have "\ z. z \ unit_disc \ poincare_distance x z = d" proof (rule wlog_zero) show "\ z. z \ unit_disc \ poincare_distance 0\<^sub>h z = d" using ex_x_axis_poincare_distance_negative[of d] \d \ 0\ by blast next show "x \ unit_disc" by fact next fix a u assume "u \ unit_disc" "cmod a < 1" assume "\z. z \ unit_disc \ poincare_distance (moebius_pt (blaschke a) u) z = d" then obtain z where *: "z \ unit_disc" "poincare_distance (moebius_pt (blaschke a) u) z = d" by auto obtain z' where z': "z = moebius_pt (blaschke a) z'" "z' \ unit_disc" using \z \ unit_disc\ using unit_disc_fix_iff[of "blaschke a"] \cmod a < 1\ using blaschke_unit_disc_fix[of a] by blast show "\z. z \ unit_disc \ poincare_distance u z = d" using * z' \u : unit_disc\ using blaschke_unit_disc_fix[of a] \cmod a < 1\ by (rule_tac x=z' in exI, simp) qed thus ?thesis using \x = y\ unfolding poincare_between_def by auto next case False show ?thesis proof (rule wlog_positive_x_axis[where P="\ y x. ?P x y"]) fix x assume "is_real x" "0 < Re x" "Re x < 1" then obtain z where z: "is_real z" "Re z \ 0" "- 1 < Re z" "of_complex z \ unit_disc" "of_complex z \ unit_disc" "of_complex z \ circline_set x_axis" "poincare_distance 0\<^sub>h (of_complex z) = d" using ex_x_axis_poincare_distance_negative[of d] \d \ 0\ by auto have "poincare_between (of_complex x) 0\<^sub>h (of_complex z)" proof (cases "z = 0") case True thus ?thesis unfolding poincare_between_def by auto next case False have "x \ 0" using \is_real x\ \Re x > 0\ by auto thus ?thesis using poincare_between_x_axis_u0v[of x z] using z \is_real x\ \x \ 0\ \Re x > 0\ False using complex_eq_if_Re_eq mult_pos_neg by fastforce qed thus "?P (of_complex x) 0\<^sub>h" using \poincare_distance 0\<^sub>h (of_complex z) = d\ \of_complex z \ unit_disc\ by blast next show "x \ unit_disc" "y \ unit_disc" by fact+ next show "y \ x" using \x \ y\ by simp next fix M u v assume "unit_disc_fix M" "u \ unit_disc" "v \ unit_disc" "u \ v" assume "?P (moebius_pt M v) (moebius_pt M u)" then obtain z where *: "z \ unit_disc" "poincare_between (moebius_pt M v) (moebius_pt M u) z" "poincare_distance (moebius_pt M u) z = d" by auto obtain z' where z': "z = moebius_pt M z'" "z' \ unit_disc" using \z \ unit_disc\ using unit_disc_fix_iff[of M] \unit_disc_fix M\ by blast thus "?P v u" using * \u \ unit_disc\ \v \ unit_disc\ \unit_disc_fix M\ by auto qed qed thus ?thesis using assms d by auto qed (* ------------------------------------------------------------------ *) subsection\Five segment axiom\ (* ------------------------------------------------------------------ *) lemma five_segment_axiom: assumes in_disc: "x \ unit_disc" "y \ unit_disc" "z \ unit_disc" "u \ unit_disc" and in_disc': "x' \ unit_disc" "y' \ unit_disc" "z' \ unit_disc" "u' \ unit_disc" and "x \ y" and betw: "poincare_between x y z" "poincare_between x' y' z'" and xy: "poincare_distance x y = poincare_distance x' y'" and xu: "poincare_distance x u = poincare_distance x' u'" and yu: "poincare_distance y u = poincare_distance y' u'" and yz: "poincare_distance y z = poincare_distance y' z'" shows "poincare_distance z u = poincare_distance z' u'" proof- from assms obtain M where M: "unit_disc_fix_f M" "M x = x'" "M u = u'" "M y = y'" using unit_disc_fix_f_congruent_triangles[of x y u] by blast have "M z = z'" proof (rule unique_poincare_distance_on_ray[where u=x' and v=y' and y="M z" and z=z' and d="poincare_distance x z"]) show "0 \ poincare_distance x z" using poincare_distance_ge0 in_disc by simp next show "x' \ y'" using M \x \ y\ using in_disc in_disc' poincare_distance_eq_0_iff xy by auto next show "poincare_distance x' (M z) = poincare_distance x z" using M in_disc unfolding unit_disc_fix_f_def by auto next show "M z \ unit_disc" using M in_disc unfolding unit_disc_fix_f_def by auto next show "poincare_distance x' z' = poincare_distance x z" using xy yz betw using poincare_between_sum_distances[of x y z] using poincare_between_sum_distances[of x' y' z'] using in_disc in_disc' by auto next show "poincare_between x' y' (M z)" using M using in_disc betw unfolding unit_disc_fix_f_def by auto qed fact+ thus ?thesis using \unit_disc_fix_f M\ using in_disc in_disc' \M u = u'\ unfolding unit_disc_fix_f_def by auto qed (* ------------------------------------------------------------------ *) subsection\Upper dimension axiom\ (* ------------------------------------------------------------------ *) lemma upper_dimension_axiom: assumes in_disc: "x \ unit_disc" "y \ unit_disc" "z \ unit_disc" "u \ unit_disc" "v \ unit_disc" assumes "poincare_distance x u = poincare_distance x v" "poincare_distance y u = poincare_distance y v" "poincare_distance z u = poincare_distance z v" "u \ v" shows "poincare_between x y z \ poincare_between y z x \ poincare_between z x y" proof (cases "x = y \ y = z \ x = z") case True thus ?thesis using in_disc by auto next case False hence "x \ y" "x \ z" "y \ z" by auto let ?cong = "\ a b a' b'. poincare_distance a b = poincare_distance a' b'" have "\ z u v. z \ unit_disc \ u \ unit_disc \ v \ unit_disc \ ?cong x u x v \ ?cong y u y v \ ?cong z u z v \ u \ v \ poincare_collinear {x, y, z}" (is "?P x y") proof (rule wlog_positive_x_axis[where P="?P"]) fix x assume x: "is_real x" "0 < Re x" "Re x < 1" hence "x \ 0" by auto have "0\<^sub>h \ circline_set x_axis" by simp show "?P 0\<^sub>h (of_complex x)" proof safe fix z u v assume in_disc: "z \ unit_disc" "u \ unit_disc" "v \ unit_disc" then obtain z' u' v' where zuv: "z = of_complex z'" "u = of_complex u'" "v = of_complex v'" using inf_or_of_complex[of z] inf_or_of_complex[of u] inf_or_of_complex[of v] by auto assume cong: "?cong 0\<^sub>h u 0\<^sub>h v" "?cong (of_complex x) u (of_complex x) v" "?cong z u z v" "u \ v" let ?r0 = "poincare_distance 0\<^sub>h u" and ?rx = "poincare_distance (of_complex x) u" have "?r0 > 0" "?rx > 0" using in_disc cong using poincare_distance_eq_0_iff[of "0\<^sub>h" u] poincare_distance_ge0[of "0\<^sub>h" u] using poincare_distance_eq_0_iff[of "0\<^sub>h" v] poincare_distance_ge0[of "0\<^sub>h" v] using poincare_distance_eq_0_iff[of "of_complex x" u] poincare_distance_ge0[of "of_complex x" u] using poincare_distance_eq_0_iff[of "of_complex x" v] poincare_distance_ge0[of "of_complex x" v] using x by (auto simp add: cmod_eq_Re) let ?pc0 = "poincare_circle 0\<^sub>h ?r0" and ?pcx = "poincare_circle (of_complex x) ?rx" have "u \ ?pc0 \ ?pcx" "v \ ?pc0 \ ?pcx" using in_disc cong by (auto simp add: poincare_circle_def) hence "u = conjugate v" using intersect_poincare_circles_x_axis[of 0 x ?r0 ?rx u v] using x \x \ 0\ \u \ v\ \?r0 > 0\ \?rx > 0\ by simp let ?ru = "poincare_distance u z" have "?ru > 0" using poincare_distance_ge0[of u z] in_disc using cong using poincare_distance_eq_0_iff[of z u] poincare_distance_eq_0_iff[of z v] using poincare_distance_eq_0_iff by force have "z \ poincare_circle u ?ru \ poincare_circle v ?ru" using cong in_disc unfolding poincare_circle_def by (simp add: poincare_distance_sym) hence "is_real z'" using intersect_poincare_circles_conjugate_centers[of u v ?ru z] \u = conjugate v\ zuv using in_disc \u \ v\ \?ru > 0\ by simp thus "poincare_collinear {0\<^sub>h, of_complex x, z}" using poincare_line_0_real_is_x_axis[of "of_complex x"] x \x \ 0\ zuv \0\<^sub>h \ circline_set x_axis\ unfolding poincare_collinear_def by (rule_tac x=x_axis in exI, auto simp add: circline_set_x_axis) qed next fix M x y assume 1: "unit_disc_fix M" "x \ unit_disc" "y \ unit_disc" "x \ y" assume 2: "?P (moebius_pt M x) (moebius_pt M y)" show "?P x y" proof safe fix z u v assume "z \ unit_disc" "u \ unit_disc" "v \ unit_disc" "?cong x u x v" "?cong y u y v" "?cong z u z v" "u \ v" hence "poincare_collinear {moebius_pt M x, moebius_pt M y, moebius_pt M z}" using 1 2[rule_format, of "moebius_pt M z" "moebius_pt M u" "moebius_pt M v"] by simp then obtain p where "is_poincare_line p" "{moebius_pt M x, moebius_pt M y, moebius_pt M z} \ circline_set p" unfolding poincare_collinear_def by auto thus "poincare_collinear {x, y, z}" using \unit_disc_fix M\ unfolding poincare_collinear_def by (rule_tac x="moebius_circline (-M) p" in exI, auto) qed qed fact+ thus ?thesis using assms using poincare_collinear3_between[of x y z] using poincare_between_rev by auto qed (* ------------------------------------------------------------------ *) subsection\Lower dimension axiom\ (* ------------------------------------------------------------------ *) lemma lower_dimension_axiom: shows "\ a \ unit_disc. \ b \ unit_disc. \ c \ unit_disc. \ poincare_between a b c \ \ poincare_between b c a \ \ poincare_between c a b" proof- let ?u = "of_complex (1/2)" and ?v = "of_complex (\/2)" have 1: "0\<^sub>h \ unit_disc" and 2: "?u \ unit_disc" and 3: "?v \ unit_disc" by simp_all have *: "\ poincare_collinear {0\<^sub>h, ?u, ?v}" proof (rule ccontr) assume "\ ?thesis" then obtain p where "is_poincare_line p" "{0\<^sub>h, ?u, ?v} \ circline_set p" unfolding poincare_collinear_def by auto moreover have "of_complex (1 / 2) \ of_complex (\ / 2)" using of_complex_inj by fastforce ultimately have "0\<^sub>h \ circline_set (poincare_line ?u ?v)" using unique_poincare_line[of ?u ?v p] by auto thus False unfolding circline_set_def by simp (transfer, transfer, simp add: vec_cnj_def) qed show ?thesis apply (rule_tac x="0\<^sub>h" in bexI, rule_tac x="?u" in bexI, rule_tac x="?v" in bexI) apply (rule ccontr, auto) using * using poincare_between_poincare_collinear[OF 1 2 3] using poincare_between_poincare_collinear[OF 2 3 1] using poincare_between_poincare_collinear[OF 3 1 2] by (metis insert_commute)+ qed (* ------------------------------------------------------------------ *) subsection\Negated Euclidean axiom\ (* ------------------------------------------------------------------ *) lemma negated_euclidean_axiom_aux: assumes "on_circline H (of_complex (1/2 + \/2))" and "is_poincare_line H" assumes "intersects_x_axis_positive H" shows "\ intersects_y_axis_positive H" using assms proof (transfer, transfer) fix H assume hh: "hermitean H \ H \ mat_zero" "is_poincare_line_cmat H" obtain A B C D where "H = (A, B, C, D)" by (cases H, auto) hence *: "is_real A" "H = (A, B, cnj B, A)" "(cmod B)\<^sup>2 > (cmod A)\<^sup>2" using hermitean_elems[of A B C D] hh by auto assume "intersects_x_axis_positive_cmat H" hence "Re A \ 0" "Re B / Re A < - 1" using * by auto assume "on_circline_cmat_cvec H (of_complex_cvec (1 / 2 + \ / 2))" hence "6*A + 4*Re B + 4*Im B = 0" using * unfolding of_real_mult apply (subst Re_express_cnj[of B]) apply (subst Im_express_cnj[of B]) apply (simp add: vec_cnj_def) apply (simp add: field_simps) done hence "Re (6*A + 4*Re B + 4*Im B) = 0" by simp hence "3*Re A + 2*Re B + 2*Im B = 0" using \is_real A\ by simp hence "3/2 + Re B/Re A + Im B/Re A = 0" using \Re A \ 0\ by (simp add: field_simps) hence "-Im B/Re A - 3/2 < -1" using \Re B / Re A < -1\ by simp hence "Im B/Re A > -1/2" by (simp add: field_simps) thus "\ intersects_y_axis_positive_cmat H" using * by simp qed lemma negated_euclidean_axiom: shows "\ a b c d t. a \ unit_disc \ b \ unit_disc \ c \ unit_disc \ d \ unit_disc \ t \ unit_disc \ poincare_between a d t \ poincare_between b d c \ a \ d \ (\ x y. x \ unit_disc \ y \ unit_disc \ poincare_between a b x \ poincare_between x t y \ \ poincare_between a c y)" proof- let ?a = "0\<^sub>h" let ?b = "of_complex (1/2)" let ?c = "of_complex (\/2)" let ?dl = "(5 - sqrt 17) / 4" let ?d = "of_complex (?dl + \*?dl)" let ?t = "of_complex (1/2 + \/2)" have "?dl \ 0" proof- have "(sqrt 17)\<^sup>2 \ 5\<^sup>2" by simp hence "sqrt 17 \ 5" by force thus ?thesis by simp qed have "?d \ ?a" proof (rule ccontr) assume "\ ?thesis" hence "?dl + \*?dl = 0" by simp hence "Re (?dl + \*?dl) = 0" by simp thus False using \?dl \ 0\ by simp qed have "?dl > 0" proof- have "(sqrt 17)\<^sup>2 < 5\<^sup>2" by (simp add: power2_eq_square) hence "sqrt 17 < 5" by (rule power2_less_imp_less, simp) thus ?thesis by simp qed have "?a \ ?b" by (metis divide_eq_0_iff of_complex_zero_iff zero_neq_numeral zero_neq_one) have "?a \ ?c" by (metis complex_i_not_zero divide_eq_0_iff of_complex_zero_iff zero_neq_numeral) show ?thesis proof (rule_tac x="?a" in exI, rule_tac x="?b" in exI, rule_tac x="?c" in exI, rule_tac x="?d" in exI, rule_tac x="?t" in exI, safe) show "?a \ unit_disc" "?b \ unit_disc" "?c \ unit_disc" "?t \ unit_disc" by (auto simp add: cmod_def power2_eq_square) have cmod_d: "cmod (?dl + \*?dl) = ?dl * sqrt 2" using \?dl > 0\ unfolding cmod_def by (simp add: real_sqrt_mult) show "?d \ unit_disc" proof- have "?dl < 1 / sqrt 2" proof- have "17\<^sup>2 < (5 * sqrt 17)\<^sup>2" by (simp add: field_simps) hence "17 < 5 * sqrt 17" by (rule power2_less_imp_less, simp) hence "?dl\<^sup>2 < (1 / sqrt 2)\<^sup>2" by (simp add: power2_eq_square field_simps) thus "?dl < 1 / sqrt 2" by (rule power2_less_imp_less, simp) qed thus ?thesis using cmod_d by (simp add: field_simps) qed have cmod_d: "1 - (cmod (to_complex ?d))\<^sup>2 = (-17 + 5*sqrt 17) / 4" (is "_ = ?cmod_d") apply (simp only: to_complex_of_complex) apply (subst cmod_d) apply (simp add: power_mult_distrib) apply (simp add: power2_eq_square field_simps) done have cmod_d_c: "(cmod (to_complex ?d - to_complex ?c))\<^sup>2 = (17 - 4*sqrt 17) / 4" (is "_ = ?cmod_dc") unfolding cmod_square by (simp add: field_simps) have cmod_c: "1 - (cmod (to_complex ?c))\<^sup>2 = 3/4" (is "_ = ?cmod_c") by (simp add: power2_eq_square) have xx: "\ x::real. x + x = 2*x" by simp have "cmod ((to_complex ?b) - (to_complex ?d)) = cmod ((to_complex ?d) - (to_complex ?c))" by (simp add: cmod_def power2_eq_square field_simps) moreover have "cmod (to_complex ?b) = cmod (to_complex ?c)" by simp ultimately have *: "poincare_distance_formula' (to_complex ?b) (to_complex ?d) = poincare_distance_formula' (to_complex ?d) (to_complex ?c)" unfolding poincare_distance_formula'_def by simp have **: "poincare_distance_formula' (to_complex ?d) (to_complex ?c) = (sqrt 17) / 3" unfolding poincare_distance_formula'_def proof (subst cmod_d, subst cmod_c, subst cmod_d_c) have "(sqrt 17 * 15)\<^sup>2 \ 51\<^sup>2" by simp hence "sqrt 17 * 15 \ 51" by force hence "sqrt 17 * 15 - 51 \ 0" by simp have "(5 * sqrt 17)\<^sup>2 \ 17\<^sup>2" by simp hence "5 * sqrt 17 \ 17" by force hence "?cmod_d * ?cmod_c \ 0" by simp hence "1 + 2 * (?cmod_dc / (?cmod_d * ?cmod_c)) = (?cmod_d * ?cmod_c + 2 * ?cmod_dc) / (?cmod_d * ?cmod_c)" using add_frac_num[of "?cmod_d * ?cmod_c" "2 * ?cmod_dc" 1] by (simp add: field_simps) also have "... = (64 * (85 - sqrt 17 * 17)) / (64 * (sqrt 17 * 15 - 51))" by (simp add: field_simps) also have "... = (85 - sqrt 17 * 17) / (sqrt 17 * 15 - 51)" by (rule mult_divide_mult_cancel_left, simp) also have "... = sqrt 17 / 3" by (subst frac_eq_eq, fact, simp, simp add: field_simps) finally show "1 + 2 * (?cmod_dc / (?cmod_d * ?cmod_c)) = sqrt 17 / 3" . qed have "sqrt 17 \ 3" proof- have "(sqrt 17)\<^sup>2 \ 3\<^sup>2" by simp thus ?thesis by (rule power2_le_imp_le, simp) qed thus "poincare_between ?b ?d ?c" unfolding poincare_between_sum_distances[OF \?b \ unit_disc\ \?d \ unit_disc\ \?c \ unit_disc\] unfolding poincare_distance_formula[OF \?b \ unit_disc\ \?d \ unit_disc\] unfolding poincare_distance_formula[OF \?d \ unit_disc\ \?c \ unit_disc\] unfolding poincare_distance_formula[OF \?b \ unit_disc\ \?c \ unit_disc\] unfolding poincare_distance_formula_def apply (subst *, subst xx, subst **, subst arcosh_double) apply (simp_all add: cmod_def power2_eq_square) done show "poincare_between ?a ?d ?t" proof (subst poincare_between_0uv[OF \?d \ unit_disc\ \?t \ unit_disc\ \?d \ ?a\]) show "?t \ 0\<^sub>h" proof (rule ccontr) assume "\ ?thesis" hence "1/2 + \/2 = 0" by simp hence "Re (1/2 + \/2) = 0" by simp thus False by simp qed next have "19\<^sup>2 \ (5 * sqrt 17)\<^sup>2" by simp hence "19 \ 5 * sqrt 17" by (rule power2_le_imp_le, simp) hence "cmod (to_complex ?d) \ cmod (to_complex ?t)" by (simp add: Let_def cmod_def power2_eq_square field_simps) moreover have "arg (to_complex ?d) = arg (to_complex ?t)" proof- have 1: "to_complex ?d = ((5 - sqrt 17) / 4) * (1 + \)" by (simp add: field_simps) have 2: "to_complex ?t = (cor (1/2)) * (1 + \)" by (simp add: field_simps) have "(sqrt 17)\<^sup>2 < 5\<^sup>2" by simp hence "sqrt 17 < 5" by (rule power2_less_imp_less, simp) hence 3: "(5 - sqrt 17) / 4 > 0" by simp have 4: "(1::real) / 2 > 0" by simp show ?thesis apply (subst 1, subst 2) apply (subst arg_mult_real_positive[OF 3]) apply (subst arg_mult_real_positive[OF 4]) by simp qed ultimately show "let d' = to_complex ?d; t' = to_complex ?t in arg d' = arg t' \ cmod d' \ cmod t'" by simp qed show "?a = ?d \ False" using \?d \ ?a\ by simp fix x y assume "x \ unit_disc" "y \ unit_disc" assume abx: "poincare_between ?a ?b x" hence "x \ circline_set x_axis" using poincare_between_poincare_line_uvz[of ?a ?b x] \x \ unit_disc\ \?a \ ?b\ using poincare_line_0_real_is_x_axis[of ?b] by (auto simp add: circline_set_x_axis) have "x \ 0\<^sub>h" using abx poincare_between_sandwich[of ?a ?b] \?a \ ?b\ by auto have "x \ positive_x_axis" using \x \ circline_set x_axis\ \x \ 0\<^sub>h\ \x \ unit_disc\ using abx poincare_between_x_axis_0uv[of "1/2" "Re (to_complex x)"] unfolding circline_set_x_axis positive_x_axis_def by (auto simp add: cmod_eq_Re abs_less_iff complex_eq_if_Re_eq) assume acy: "poincare_between ?a ?c y" hence "y \ circline_set y_axis" using poincare_between_poincare_line_uvz[of ?a ?c y] \y \ unit_disc\ \?a \ ?c\ using poincare_line_0_imag_is_y_axis[of ?c] by (auto simp add: circline_set_y_axis) have "y \ 0\<^sub>h" using acy poincare_between_sandwich[of ?a ?c] \?a \ ?c\ by auto have "y \ positive_y_axis" proof- have " \x. \x \ 0; poincare_between 0\<^sub>h (of_complex (\ / 2)) (of_complex x); is_imag x; - 1 < Im x\ \ 0 < Im x" by (smt add.left_neutral complex.expand divide_complex_def complex_eq divide_less_0_1_iff divide_less_eq_1_pos imaginary_unit.simps(1) mult.left_neutral of_real_1 of_real_add of_real_divide of_real_eq_0_iff one_add_one poincare_between_y_axis_0uv zero_complex.simps(1) zero_complex.simps(2) zero_less_divide_1_iff) thus ?thesis using \y \ circline_set y_axis\ \y \ 0\<^sub>h\ \y \ unit_disc\ using acy unfolding circline_set_y_axis positive_y_axis_def by (auto simp add: cmod_eq_Im abs_less_iff) qed have "x \ y" using \x \ positive_x_axis\ \y \ positive_y_axis\ unfolding positive_x_axis_def positive_y_axis_def circline_set_x_axis circline_set_y_axis by auto assume xty: "poincare_between x ?t y" let ?xy = "poincare_line x y" have "?t \ circline_set ?xy" using xty poincare_between_poincare_line_uzv[OF \x \ y\ \x \ unit_disc\ \y \ unit_disc\ \?t \ unit_disc\] by simp moreover have "?xy \ x_axis" using poincare_line_circline_set[OF \x \ y\] \y \ positive_y_axis\ by (auto simp add: circline_set_x_axis positive_y_axis_def) hence "intersects_x_axis_positive ?xy" using intersects_x_axis_positive_iff[of "?xy"] \x \ y\ \x \ unit_disc\ \x \ positive_x_axis\ by auto moreover have "?xy \ y_axis" using poincare_line_circline_set[OF \x \ y\] \x \ positive_x_axis\ by (auto simp add: circline_set_y_axis positive_x_axis_def) hence "intersects_y_axis_positive ?xy" using intersects_y_axis_positive_iff[of "?xy"] \x \ y\ \y \ unit_disc\ \y \ positive_y_axis\ by auto ultimately show False using negated_euclidean_axiom_aux[of ?xy] \x \ y\ unfolding circline_set_def by auto qed qed text \Alternate form of the Euclidean axiom -- this one is much easier to prove\ lemma negated_euclidean_axiom': shows "\ a b c. a \ unit_disc \ b \ unit_disc \ c \ unit_disc \ \(poincare_collinear {a, b, c}) \ \(\ x. x \ unit_disc \ poincare_distance a x = poincare_distance b x \ poincare_distance a x = poincare_distance c x)" proof- let ?a = "of_complex (\/2)" let ?b = "of_complex (-\/2)" let ?c = "of_complex (1/5)" have "(\/2) \ (-\/2)" by simp hence "?a \ ?b" by (metis to_complex_of_complex) have "(\/2) \ (1/5)" by simp hence "?a \ ?c" by (metis to_complex_of_complex) have "(-\/2) \ (1/5)" by (simp add: minus_equation_iff) hence "?b \ ?c" by (metis to_complex_of_complex) have "?a \ unit_disc" "?b \ unit_disc" "?c \ unit_disc" by auto moreover have "\(poincare_collinear {?a, ?b, ?c})" unfolding poincare_collinear_def proof(rule ccontr) assume " \ (\p. is_poincare_line p \ {?a, ?b, ?c} \ circline_set p)" then obtain p where "is_poincare_line p \ {?a, ?b, ?c} \ circline_set p" by auto let ?ab = "poincare_line ?a ?b" have "p = ?ab" using \is_poincare_line p \ {?a, ?b, ?c} \ circline_set p\ using unique_poincare_line[of ?a ?b] \?a \ ?b\ \?a \ unit_disc\ \?b \ unit_disc\ by auto have "?c \ circline_set ?ab" proof(rule ccontr) assume "\ ?c \ circline_set ?ab" have "poincare_between ?a 0\<^sub>h ?b" unfolding poincare_between_def using cross_ratio_0inf by auto hence "0\<^sub>h \ circline_set ?ab" using \?a \ ?b\ \?a \ unit_disc\ \?b \ unit_disc\ using poincare_between_poincare_line_uzv zero_in_unit_disc by blast hence "?ab = poincare_line 0\<^sub>h ?a" using unique_poincare_line[of ?a ?b] \?a \ ?b\ \?a \ unit_disc\ \?b \ unit_disc\ using \is_poincare_line p \ {?a, ?b, ?c} \ circline_set p\ using \p = ?ab\ poincare_line_circline_set(1) unique_poincare_line by (metis add.inverse_neutral divide_minus_left of_complex_zero_iff zero_in_unit_disc) hence "(\/2) * cnj(1/5) = cnj(\/2) * (1/5)" using poincare_collinear_zero_iff[of "(\/2)" "(1/5)"] using \?a \ ?c\ \\ ?c \ circline_set ?ab\ \?a \ unit_disc\ \?c \ unit_disc\ \p = ?ab\ using \0\<^sub>h \ circline_set ?ab\ \is_poincare_line p \ {?a, ?b, ?c} \ circline_set p\ using poincare_collinear_def by auto thus False by simp qed thus False using \p = ?ab\ \is_poincare_line p \ {?a, ?b, ?c} \ circline_set p\ by auto qed moreover have "\(\ x. x \ unit_disc \ poincare_distance ?a x = poincare_distance ?b x \ poincare_distance ?a x = poincare_distance ?c x)" proof(rule ccontr) assume "\ ?thesis" then obtain x where "x \ unit_disc" "poincare_distance ?a x = poincare_distance ?b x" "poincare_distance ?a x = poincare_distance ?c x" by blast let ?x = "to_complex x" have "poincare_distance_formula' (\/2) ?x = poincare_distance_formula' (-\/2) ?x" using \poincare_distance ?a x = poincare_distance ?b x\ using \x \ unit_disc\ \?a \ unit_disc\ \?b \ unit_disc\ by (metis cosh_dist to_complex_of_complex) hence "(cmod (\ / 2 - ?x))\<^sup>2 = (cmod (- \ / 2 - ?x))\<^sup>2" unfolding poincare_distance_formula'_def apply (simp add:field_simps) using \x \ unit_disc\ unit_disc_cmod_square_lt_1 by fastforce hence "Im ?x = 0" unfolding cmod_def by (simp add: power2_eq_iff) have "1 - (Re ?x)\<^sup>2 \ 0" using \x \ unit_disc\ unit_disc_cmod_square_lt_1 using cmod_power2 by force hence "24 - 24 * (Re ?x)\<^sup>2 \ 0" by simp have "poincare_distance_formula' (\/2) ?x = poincare_distance_formula' (1/5) ?x" using \poincare_distance ?a x = poincare_distance ?c x\ using \x \ unit_disc\ \?a \ unit_disc\ \?c \ unit_disc\ by (metis cosh_dist to_complex_of_complex) hence "(2 + 8 * (Re ?x)\<^sup>2) /(3 - 3 * (Re ?x)\<^sup>2) = 2 * (1 - Re ?x * 5)\<^sup>2 / (24 - 24 * (Re ?x)\<^sup>2)" (is "?lhs = ?rhs") unfolding poincare_distance_formula'_def apply (simp add:field_simps) unfolding cmod_def using \Im ?x = 0\ by (simp add:field_simps) hence *: "?lhs * (24 - 24 * (Re ?x)\<^sup>2) = ?rhs * (24 - 24 * (Re ?x)\<^sup>2) " using \(24 - 24 * (Re ?x)\<^sup>2) \ 0\ by simp have "?lhs * (24 - 24 * (Re ?x)\<^sup>2) = (2 + 8 * (Re ?x)\<^sup>2) * 8" using \(24 - 24 * (Re ?x)\<^sup>2) \ 0\ \1 - (Re ?x)\<^sup>2 \ 0\ by (simp add:field_simps) have "?rhs * (24 - 24 * (Re ?x)\<^sup>2) = 2 * (1 - Re ?x * 5)\<^sup>2" using \(24 - 24 * (Re ?x)\<^sup>2) \ 0\ \1 - (Re ?x)\<^sup>2 \ 0\ by (simp add:field_simps) hence "(2 + 8 * (Re ?x)\<^sup>2) * 8 = 2 * (1 - Re ?x * 5)\<^sup>2" using * \?lhs * (24 - 24 * (Re ?x)\<^sup>2) = (2 + 8 * (Re ?x)\<^sup>2) * 8\ by simp hence "7 * (Re ?x)\<^sup>2 + 10 * (Re ?x) + 7 = 0" by (simp add:field_simps comm_ring_1_class.power2_diff) thus False using discriminant_iff[of 7 "Re (to_complex x)" 10 7] discrim_def[of 7 10 7] by auto qed ultimately show ?thesis apply (rule_tac x="?a" in exI) apply (rule_tac x="?b" in exI) apply (rule_tac x="?c" in exI) by auto qed (* ------------------------------------------------------------------ *) subsection\Continuity axiom\ (* ------------------------------------------------------------------ *) text \The set $\phi$ is on the left of the set $\psi$\ abbreviation set_order where "set_order A \ \ \ \x\ unit_disc. \y\ unit_disc. \ x \ \ y \ poincare_between A x y" text \The point $B$ is between the sets $\phi$ and $\psi$\ abbreviation point_between_sets where "point_between_sets \ B \ \ \x\ unit_disc. \y\ unit_disc. \ x \ \ y \ poincare_between x B y" lemma continuity: assumes "\ A \ unit_disc. set_order A \ \" shows "\ B \ unit_disc. point_between_sets \ B \" proof (cases "(\ x0 \ unit_disc. \ x0) \ (\ y0 \ unit_disc. \ y0)") case False thus ?thesis using assms by blast next case True then obtain Y0 where "\ Y0" "Y0 \ unit_disc" by auto obtain A where *: "A \ unit_disc" "set_order A \ \" using assms by auto show ?thesis proof(cases "\ x \ unit_disc. \ x \ x = A") case True thus ?thesis using \A \ unit_disc\ using poincare_between_nonstrict(1) by blast next case False then obtain X0 where "\ X0" "X0 \ A" "X0 \ unit_disc" by auto have "Y0 \ A" proof(rule ccontr) assume "\ Y0 \ A" hence "\ x \ unit_disc. \ x \ poincare_between A x A" using * \\ Y0\ by (cases A) force hence "\ x \ unit_disc. \ x \ x = A" using * poincare_between_sandwich by blast thus False using False by auto qed show ?thesis proof (cases "\ B \ unit_disc. \ B \ \ B") case True then obtain B where "B \ unit_disc" "\ B" "\ B" by auto hence "\ x \ unit_disc. \ x \ poincare_between A x B" using * by auto have "\ y \ unit_disc. \ y \ poincare_between A B y" using * \B \ unit_disc\ \\ B\ by auto show ?thesis proof(rule+) show "B \ unit_disc" by fact next fix x y assume "x \ unit_disc" "y \ unit_disc" "\ x \ \ y" hence "poincare_between A x B" "poincare_between A B y" using \\ x \ unit_disc. \ x \ poincare_between A x B\ using \\ y \ unit_disc. \ y \ poincare_between A B y\ by simp+ thus "poincare_between x B y" using \x \ unit_disc\ \y \ unit_disc\ \B \ unit_disc\ \A \ unit_disc\ using poincare_between_transitivity[of A x B y] by simp qed next case False have "poincare_between A X0 Y0" using \\ X0\ \\ Y0\ * \Y0 \ unit_disc\ \X0 \ unit_disc\ by auto have "\ \. \ \. set_order A \ \ \ \ (\ B \ unit_disc. \ B \ \ B) \ \ X0 \ (\ y \ unit_disc. \ y) \ (\ x \ unit_disc. \ x) \ (\ B \ unit_disc. point_between_sets \ B \)" (is "?P A X0") proof (rule wlog_positive_x_axis[where P="?P"]) show "A \ unit_disc" by fact next show "X0 \ unit_disc" by fact next show "A \ X0" using \X0 \ A\ by simp next fix M u v let ?M = "\ x. moebius_pt M x" let ?Mu = "?M u" and ?Mv = "?M v" assume hip: "unit_disc_fix M" "u \ unit_disc" "v \ unit_disc" "u \ v" "?P ?Mu ?Mv" show "?P u v" proof safe fix \ \ x y assume "set_order u \ \" "\ (\B\unit_disc. \ B \ \ B)" "\ v" "y \ unit_disc" "\ y" "x \ unit_disc" "\ x" let ?M\ = "\ X'. \ X. \ X \ ?M X = X'" let ?M\ = "\ X'. \ X. \ X \ ?M X = X'" obtain M\ where "M\ = ?M\" by simp obtain M\ where "M\ = ?M\" by simp have "M\ ?Mv" using \\ v\ using \M\ = ?M\\ by blast moreover have "\ (\ B \unit_disc. M\ B \ M\ B)" using \\ (\B\unit_disc. \ B \ \ B)\ using \M\ = ?M\\ \M\ = ?M\\ by (metis hip(1) moebius_pt_invert unit_disc_fix_discI unit_disc_fix_moebius_inv) moreover have "\ y \ unit_disc. M\ y" using \y \ unit_disc\ \\ y\ \M\ = ?M\\ \unit_disc_fix M\ by auto moreover have "set_order ?Mu ?M\ ?M\" proof ((rule ballI)+, rule impI) fix Mx My assume "Mx \ unit_disc" "My \ unit_disc" "?M\ Mx \ ?M\ My" then obtain x y where "\ x \ ?M x = Mx" "\ y \ ?M y = My" by blast hence "x \ unit_disc" "y \ unit_disc" using \Mx \ unit_disc\ \My \ unit_disc\ \unit_disc_fix M\ by (metis moebius_pt_comp_inv_left unit_disc_fix_discI unit_disc_fix_moebius_inv)+ hence "poincare_between u x y" using \set_order u \ \\ using \Mx \ unit_disc\ \My \ unit_disc\ \\ x \ ?M x = Mx\ \\ y \ ?M y = My\ by blast then show "poincare_between ?Mu Mx My" using \\ x \ ?M x = Mx\ \\ y \ ?M y = My\ using \x \ unit_disc\ \y \ unit_disc\ \u \ unit_disc\ \unit_disc_fix M\ using unit_disc_fix_moebius_preserve_poincare_between by blast qed hence "set_order ?Mu M\ M\" using \M\ = ?M\\ \M\ = ?M\\ by simp ultimately have "\ Mb \ unit_disc. point_between_sets M\ Mb M\" using hip(5) by blast then obtain Mb where bbb: "Mb \ unit_disc" "point_between_sets ?M\ Mb ?M\" using \M\ = ?M\\ \M\ = ?M\\ by auto let ?b = "moebius_pt (moebius_inv M) Mb" show "\ b \ unit_disc. point_between_sets \ b \" proof (rule_tac x="?b" in bexI, (rule ballI)+, rule impI) fix x y assume "x \ unit_disc" "y \ unit_disc" "\ x \ \ y" hence "poincare_between u x y" using \set_order u \ \\ by blast let ?Mx = "?M x" and ?My = "?M y" have "?M\ ?Mx" "?M\ ?My" using \\ x \ \ y\ by blast+ have "?Mx \ unit_disc" "?My \ unit_disc" using \x \ unit_disc\ \unit_disc_fix M\ \y \ unit_disc\ by auto hence "poincare_between ?Mx Mb ?My" using \?M\ ?Mx\ \?M\ ?My\ \?Mx \ unit_disc\ \?My \ unit_disc\ bbb by auto then show "poincare_between x ?b y" using \unit_disc_fix M\ using \x \ unit_disc\ \y \ unit_disc\ \Mb \ unit_disc\ \?Mx \ unit_disc\ \?My \ unit_disc\ using unit_disc_fix_moebius_preserve_poincare_between[of M x ?b y] by auto next show "?b \ unit_disc" using bbb \unit_disc_fix M\ by auto qed qed next fix X assume xx: "is_real X" "0 < Re X" "Re X < 1" let ?X = "of_complex X" show "?P 0\<^sub>h ?X" proof ((rule allI)+, rule impI, (erule conjE)+) fix \ \ assume "set_order 0\<^sub>h \ \" "\ (\B\unit_disc. \ B \ \ B)" "\ ?X" "\y\unit_disc. \ y" "\x\unit_disc. \ x" have "?X \ unit_disc" using xx by (simp add: cmod_eq_Re) have \pos: "\ y \ unit_disc. \ y \ (is_real (to_complex y) \ Re (to_complex y) > 0)" proof(rule ballI, rule impI) fix y let ?y = "to_complex y" assume "y \ unit_disc" "\ y" hence "poincare_between 0\<^sub>h ?X y" using \set_order 0\<^sub>h \ \\ using \?X \ unit_disc\ \\ ?X\ by auto thus "is_real ?y \ 0 < Re ?y" using xx \?X \ unit_disc\ \y \ unit_disc\ by (metis (mono_tags, hide_lams) arg_0_iff of_complex_zero_iff poincare_between_0uv poincare_between_sandwich to_complex_of_complex unit_disc_to_complex_inj zero_in_unit_disc) qed have \noneg: "\ x \ unit_disc. \ x \ (is_real (to_complex x) \ Re (to_complex x) \ 0)" proof(rule ballI, rule impI) fix x assume "x \ unit_disc" "\ x" obtain y where "y \ unit_disc" "\ y" using \\ y \ unit_disc. \ y\ by blast let ?x = "to_complex x" and ?y = "to_complex y" have "is_real ?y" "Re ?y > 0" using \pos \\ y\ \y \ unit_disc\ by auto have "poincare_between 0\<^sub>h x y" using \set_order 0\<^sub>h \ \\ using \x \ unit_disc\ \\ x\ \y\unit_disc\ \\ y\ by auto thus "is_real ?x \ 0 \ Re ?x" using \x \ unit_disc\ \y \ unit_disc\ \is_real (to_complex y)\ \\ y\ using \set_order 0\<^sub>h \ \\ using \\ ?X\ \?X \ unit_disc\ \Re ?y > 0\ by (metis arg_0_iff le_less of_complex_zero poincare_between_0uv to_complex_of_complex zero_complex.simps(1) zero_complex.simps(2)) qed have \less\: "\x\unit_disc. \y\unit_disc. \ x \ \ y \ Re (to_complex x) < Re (to_complex y)" proof((rule ballI)+, rule impI) fix x y let ?x = "to_complex x" and ?y = "to_complex y" assume "x \ unit_disc" "y \ unit_disc" "\ x \ \ y" hence "poincare_between 0\<^sub>h x y" using \set_order 0\<^sub>h \ \\ by auto moreover have "is_real ?x" "Re ?x \ 0" using \noneg using \x \ unit_disc\ \\ x \ \ y\ by auto moreover have "is_real ?y" "Re ?y > 0" using \pos using \y \ unit_disc\ \\ x \ \ y\ by auto ultimately have "Re ?x \ Re ?y" using \x \ unit_disc\ \y \ unit_disc\ - by (metis Re_complex_of_real arg_0_iff le_less of_complex_zero poincare_between_0uv rcis_cmod_arg rcis_zero_arg to_complex_of_complex) + by (metis Re_complex_of_real arg_0_iff le_less of_complex_zero poincare_between_0uv rcis_cmod_Arg rcis_zero_arg to_complex_of_complex) have "Re ?x \ Re ?y" using \\ x \ \ y\ \is_real ?x\ \is_real ?y\ using \\ (\B\unit_disc. \ B \ \ B)\ \x \ unit_disc\ \y \ unit_disc\ by (metis complex.expand unit_disc_to_complex_inj) thus "Re ?x < Re ?y" using \Re ?x \ Re ?y\ by auto qed have "\ b \ unit_disc. \ x \ unit_disc. \ y \ unit_disc. is_real (to_complex b) \ (\ x \ \ y \ (Re (to_complex x) \ Re (to_complex b) \ Re (to_complex b) \ Re (to_complex y)))" proof- let ?Phi = "{x. (of_complex (cor x)) \ unit_disc \ \ (of_complex (cor x))}" have "\ x \ unit_disc. \ x \ Re (to_complex x) \ Sup ?Phi" proof(safe) fix x let ?x = "to_complex x" assume "x \ unit_disc" "\ x" hence "is_real ?x" "Re ?x \ 0" using \noneg by auto hence "cor (Re ?x) = ?x" using complex_of_real_Re by blast hence "of_complex (cor (Re ?x)) \ unit_disc" using \x \ unit_disc\ by (metis inf_notin_unit_disc of_complex_to_complex) moreover have "\ (of_complex (cor (Re ?x)))" using \cor (Re ?x) = ?x\ \\ x\ \x \ unit_disc\ by (metis inf_notin_unit_disc of_complex_to_complex) ultimately have "Re ?x \ ?Phi" by auto have "\M. \x \ ?Phi. x \ M" using \less\ using \\ y \ unit_disc. \ y\ by (metis (mono_tags, lifting) Re_complex_of_real le_less mem_Collect_eq to_complex_of_complex) thus "Re ?x \ Sup ?Phi" using cSup_upper[of "Re ?x" ?Phi] unfolding bdd_above_def using \Re ?x \ ?Phi\ by auto qed have "\ y \ unit_disc. \ y \ Sup ?Phi \ Re (to_complex y)" proof (safe) fix y let ?y = "to_complex y" assume "\ y" "y \ unit_disc" show "Sup ?Phi \ Re ?y" proof (rule ccontr) assume "\ ?thesis" hence "Re ?y < Sup ?Phi" by auto have "\ x. \ (of_complex (cor x)) \ (of_complex (cor x)) \ unit_disc" proof - obtain x' where "x' \ unit_disc" "\ x'" using \\ x \ unit_disc. \ x\ by blast let ?x' = "to_complex x'" have "is_real ?x'" using \x' \ unit_disc\ \\ x'\ using \noneg by auto hence "cor (Re ?x') = ?x'" using complex_of_real_Re by blast hence "x' = of_complex (cor (Re ?x'))" using \x' \ unit_disc\ by (metis inf_notin_unit_disc of_complex_to_complex) show ?thesis apply (rule_tac x="Re ?x'" in exI) using \x' \ unit_disc\ apply (subst (asm) \x' = of_complex (cor (Re ?x'))\, simp) using \\ x'\ by (subst (asm) (2) \x' = of_complex (cor (Re ?x'))\, simp) qed hence "?Phi \ {}" by auto then obtain x where "\ (of_complex (cor x))" "Re ?y < x" "(of_complex (cor x)) \ unit_disc" using \Re ?y < Sup ?Phi\ using less_cSupE[of "Re ?y" ?Phi] by auto moreover have "Re ?y < Re (to_complex (of_complex (cor x)))" using \Re ?y < x\ by simp ultimately show False using \less\ using \\ y\ \y \ unit_disc\ by (metis less_not_sym) qed qed thus ?thesis using \\ x \ unit_disc. \ x \ Re (to_complex x) \ Sup ?Phi\ apply (rule_tac x="(of_complex (cor (Sup ?Phi)))" in bexI, simp) using \\y\unit_disc. \ y\ \\ ?X\ \?X \ unit_disc\ using \\y\unit_disc. \ y \ is_real (to_complex y) \ 0 < Re (to_complex y)\ by (smt complex_of_real_Re inf_notin_unit_disc norm_of_real of_complex_to_complex to_complex_of_complex unit_disc_iff_cmod_lt_1 xx(2)) qed then obtain B where "B \ unit_disc" "is_real (to_complex B)" "\x\unit_disc. \y\unit_disc. \ x \ \ y \ Re (to_complex x) \ Re (to_complex B) \ Re (to_complex B) \ Re (to_complex y)" by blast show "\ b \ unit_disc. point_between_sets \ b \" proof (rule_tac x="B" in bexI) show "B \ unit_disc" by fact next show "point_between_sets \ B \" proof ((rule ballI)+, rule impI) fix x y let ?x = "to_complex x" and ?y = "to_complex y" and ?B = "to_complex B" assume "x \ unit_disc" "y \ unit_disc" "\ x \ \ y" hence "Re ?x \ Re ?B \ Re ?B \ Re ?y" using \\x\unit_disc. \y\unit_disc. \ x \ \ y \ Re (to_complex x) \ Re ?B \ Re (to_complex B) \ Re (to_complex y)\ by auto moreover have "is_real ?x" "Re ?x \ 0" using \noneg using \x \ unit_disc\ \\ x \ \ y\ by auto moreover have "is_real ?y" "Re ?y > 0" using \pos using \y \ unit_disc\ \\ x \ \ y\ by auto moreover have "cor (Re ?x) = ?x" using complex_of_real_Re \is_real ?x\ by blast hence "x = of_complex (cor (Re ?x))" using \x \ unit_disc\ by (metis inf_notin_unit_disc of_complex_to_complex) moreover have "cor (Re ?y) = ?y" using complex_of_real_Re \is_real ?y\ by blast hence "y = of_complex (cor (Re ?y))" using \y \ unit_disc\ by (metis inf_notin_unit_disc of_complex_to_complex) moreover have "cor (Re ?B) = ?B" using complex_of_real_Re \is_real (to_complex B)\ by blast hence "B = of_complex (cor (Re ?B))" using \B \ unit_disc\ by (metis inf_notin_unit_disc of_complex_to_complex) ultimately show "poincare_between x B y" using \is_real (to_complex B)\ \x \ unit_disc\ \y \ unit_disc\ \B \ unit_disc\ using poincare_between_x_axis_uvw[of "Re (to_complex x)" "Re (to_complex B)" "Re (to_complex y)"] - by (smt Re_complex_of_real arg_0_iff poincare_between_nonstrict(1) rcis_cmod_arg rcis_zero_arg unit_disc_iff_cmod_lt_1) + by (smt Re_complex_of_real arg_0_iff poincare_between_nonstrict(1) rcis_cmod_Arg rcis_zero_arg unit_disc_iff_cmod_lt_1) qed qed qed qed thus ?thesis using False \\ X0\ \\ Y0\ * \Y0 \ unit_disc\ \X0 \ unit_disc\ by auto qed qed qed (* ------------------------------------------------------------------ *) subsection\Limiting parallels axiom\ (* ------------------------------------------------------------------ *) text \Auxiliary definitions\ definition poincare_on_line where "poincare_on_line p a b \ poincare_collinear {p, a, b}" definition poincare_on_ray where "poincare_on_ray p a b \ poincare_between a p b \ poincare_between a b p" definition poincare_in_angle where "poincare_in_angle p a b c \ b \ a \ b \ c \ p \ b \ (\ x \ unit_disc. poincare_between a x c \ x \ a \ x \ c \ poincare_on_ray p b x)" definition poincare_ray_meets_line where "poincare_ray_meets_line a b c d \ (\ x \ unit_disc. poincare_on_ray x a b \ poincare_on_line x c d)" text \All points on ray are collinear\ lemma poincare_on_ray_poincare_collinear: assumes "p \ unit_disc" and "a \ unit_disc" and "b \ unit_disc" and "poincare_on_ray p a b" shows "poincare_collinear {p, a, b}" using assms poincare_between_poincare_collinear unfolding poincare_on_ray_def by (metis insert_commute) text \H-isometries preserve all defined auxiliary relations\ lemma unit_disc_fix_preserves_poincare_on_line [simp]: assumes "unit_disc_fix M" and "p \ unit_disc" "a \ unit_disc" "b \ unit_disc" shows "poincare_on_line (moebius_pt M p) (moebius_pt M a) (moebius_pt M b) \ poincare_on_line p a b" using assms unfolding poincare_on_line_def by auto lemma unit_disc_fix_preserves_poincare_on_ray [simp]: assumes "unit_disc_fix M" "p \ unit_disc" "a \ unit_disc" "b \ unit_disc" shows "poincare_on_ray (moebius_pt M p) (moebius_pt M a) (moebius_pt M b) \ poincare_on_ray p a b" using assms unfolding poincare_on_ray_def by auto lemma unit_disc_fix_preserves_poincare_in_angle [simp]: assumes "unit_disc_fix M" "p \ unit_disc" "a \ unit_disc" "b \ unit_disc" "c \ unit_disc" shows "poincare_in_angle (moebius_pt M p) (moebius_pt M a) (moebius_pt M b) (moebius_pt M c) \ poincare_in_angle p a b c" (is "?lhs \ ?rhs") proof assume "?lhs" then obtain Mx where *: "Mx \ unit_disc" "poincare_between (moebius_pt M a) Mx (moebius_pt M c)" "Mx \ moebius_pt M a" "Mx \ moebius_pt M c" "poincare_on_ray (moebius_pt M p) (moebius_pt M b) Mx" "moebius_pt M b \ moebius_pt M a" "moebius_pt M b \ moebius_pt M c" "moebius_pt M p \ moebius_pt M b" unfolding poincare_in_angle_def by auto obtain x where "Mx = moebius_pt M x" "x \ unit_disc" by (metis "*"(1) assms(1) image_iff unit_disc_fix_iff) thus ?rhs using * assms unfolding poincare_in_angle_def by auto next assume ?rhs then obtain x where *: "x \ unit_disc" "poincare_between a x c" "x \ a" "x \ c" "poincare_on_ray p b x" "b \ a" "b \ c" "p \ b" unfolding poincare_in_angle_def by auto thus ?lhs using assms unfolding poincare_in_angle_def by auto (rule_tac x="moebius_pt M x" in bexI, auto) qed lemma unit_disc_fix_preserves_poincare_ray_meets_line [simp]: assumes "unit_disc_fix M" "a \ unit_disc" "b \ unit_disc" "c \ unit_disc" "d \ unit_disc" shows "poincare_ray_meets_line (moebius_pt M a) (moebius_pt M b) (moebius_pt M c) (moebius_pt M d) \ poincare_ray_meets_line a b c d" (is "?lhs \ ?rhs") proof assume ?lhs then obtain Mx where *: "Mx \ unit_disc" "poincare_on_ray Mx (moebius_pt M a) (moebius_pt M b)" "poincare_on_line Mx (moebius_pt M c) (moebius_pt M d)" unfolding poincare_ray_meets_line_def by auto obtain x where "Mx = moebius_pt M x" "x \ unit_disc" by (metis "*"(1) assms(1) image_iff unit_disc_fix_iff) thus ?rhs using assms * unfolding poincare_ray_meets_line_def poincare_on_line_def by auto next assume ?rhs then obtain x where *: "x \ unit_disc" "poincare_on_ray x a b" "poincare_on_line x c d" unfolding poincare_ray_meets_line_def by auto thus ?lhs using assms * unfolding poincare_ray_meets_line_def poincare_on_line_def by auto (rule_tac x="moebius_pt M x" in bexI, auto) qed text \H-lines that intersect on the absolute do not meet (they do not share a common h-point)\ lemma tangent_not_meet: assumes "x1 \ unit_disc" and "x2 \ unit_disc" and "x1 \ x2" and "\ poincare_collinear {0\<^sub>h, x1, x2}" assumes "i \ ideal_points (poincare_line x1 x2)" "a \ unit_disc" "a \ 0\<^sub>h" "poincare_collinear {0\<^sub>h, a, i}" shows "\ poincare_ray_meets_line 0\<^sub>h a x1 x2" proof (rule ccontr) assume "\ ?thesis" then obtain x where "x \ unit_disc" "poincare_on_ray x 0\<^sub>h a" "poincare_collinear {x, x1, x2}" unfolding poincare_ray_meets_line_def poincare_on_line_def by auto have "poincare_collinear {0\<^sub>h, a, x}" using \poincare_on_ray x 0\<^sub>h a\ \x \ unit_disc\ \a \ unit_disc\ by (meson poincare_between_poincare_collinear poincare_between_rev poincare_on_ray_def poincare_on_ray_poincare_collinear zero_in_unit_disc) have "x \ 0\<^sub>h" using \\ poincare_collinear {0\<^sub>h, x1, x2}\ \poincare_collinear {x, x1, x2}\ unfolding poincare_collinear_def by (auto simp add: assms(2) assms(3) poincare_between_rev) let ?l1 = "poincare_line 0\<^sub>h a" let ?l2 = "poincare_line x1 x2" have "i \ circline_set unit_circle" using \i \ ideal_points (poincare_line x1 x2)\ using assms(3) ideal_points_on_unit_circle is_poincare_line_poincare_line by blast have "i \ circline_set ?l1" using \poincare_collinear {0\<^sub>h, a, i}\ unfolding poincare_collinear_def using \a \ unit_disc\ \a \ 0\<^sub>h\ by (metis insert_subset unique_poincare_line zero_in_unit_disc) moreover have "x \ circline_set ?l1" using \a \ unit_disc\ \a \ 0\<^sub>h\ \poincare_collinear {0\<^sub>h, a, x}\ \x \ unit_disc\ by (metis poincare_collinear3_between poincare_between_poincare_line_uvz poincare_between_poincare_line_uzv poincare_line_sym zero_in_unit_disc) moreover have "inversion x \ circline_set ?l1" using \poincare_collinear {0\<^sub>h, a, x}\ using poincare_line_inversion_full[of "0\<^sub>h" a x] \a \ unit_disc\ \a \ 0\<^sub>h\ \x \ unit_disc\ by (metis poincare_collinear3_between is_poincare_line_inverse_point is_poincare_line_poincare_line poincare_between_poincare_line_uvz poincare_between_poincare_line_uzv poincare_line_sym zero_in_unit_disc) moreover have "x \ circline_set ?l2" using \poincare_collinear {x, x1, x2}\ \x1 \ x2\ \x1 \ unit_disc\ \x2 \ unit_disc\ \x \ unit_disc\ by (metis insert_commute inversion_noteq_unit_disc poincare_between_poincare_line_uvz poincare_between_poincare_line_uzv poincare_collinear3_iff poincare_line_sym_general) moreover hence "inversion x \ circline_set ?l2" using \x1 \ x2\ \x1 \ unit_disc\ \x2 \ unit_disc\ \x \ unit_disc\ using poincare_line_inversion_full[of x1 x2 x] unfolding circline_set_def by auto moreover have "i \ circline_set ?l2" using \x1 \ x2\ \x1 \ unit_disc\ \x2 \ unit_disc\ using \i \ ideal_points ?l2\ by (simp add: ideal_points_on_circline) moreover have "x \ inversion x" using \x \ unit_disc\ using inversion_noteq_unit_disc by fastforce moreover have "x \ i" using \x \ unit_disc\ using \i \ circline_set unit_circle\ circline_set_def inversion_noteq_unit_disc by fastforce+ moreover have "inversion x \ i" using \i \ circline_set unit_circle\ \x \ i\ circline_set_def inversion_unit_circle by fastforce ultimately have "?l1 = ?l2" using unique_circline_set[of x "inversion x" i] by blast hence "0\<^sub>h \ circline_set ?l2" by (metis \a \ 0\<^sub>h\ poincare_line_circline_set(1)) thus False using \\ poincare_collinear {0\<^sub>h, x1, x2}\ unfolding poincare_collinear_def using \poincare_collinear {x, x1, x2}\ \x1 \ x2\ \x1 \ unit_disc\ \x2 \ unit_disc\ poincare_collinear_def unique_poincare_line by auto qed lemma limiting_parallels: assumes "a \ unit_disc" and "x1 \ unit_disc" and "x2 \ unit_disc" and "\ poincare_on_line a x1 x2" shows "\a1\unit_disc. \a2\unit_disc. \ poincare_on_line a a1 a2 \ \ poincare_ray_meets_line a a1 x1 x2 \ \ poincare_ray_meets_line a a2 x1 x2 \ (\a'\unit_disc. poincare_in_angle a' a1 a a2 \ poincare_ray_meets_line a a' x1 x2)" (is "?P a x1 x2") proof- have "\ poincare_collinear {a, x1, x2}" using \\ poincare_on_line a x1 x2\ unfolding poincare_on_line_def by simp have "\ x1 x2. x1 \ unit_disc \ x2 \ unit_disc \ \ poincare_collinear {a, x1, x2} \ ?P a x1 x2" (is "?Q a") proof (rule wlog_zero[OF \a \ unit_disc\]) fix a u assume *: "u \ unit_disc" "cmod a < 1" hence uf: "unit_disc_fix (blaschke a)" by simp assume **: "?Q (moebius_pt (blaschke a) u)" show "?Q u" proof safe fix x1 x2 let ?M = "moebius_pt (blaschke a)" assume xx: "x1 \ unit_disc" "x2 \ unit_disc" "\ poincare_collinear {u, x1, x2}" hence MM: "?M x1 \ unit_disc \ ?M x2\ unit_disc \ \ poincare_collinear {?M u, ?M x1, ?M x2}" using * by auto show "?P u x1 x2" (is "\a1\unit_disc. \a2\unit_disc. ?P' a1 a2 u x1 x2") proof- obtain Ma1 Ma2 where MM: "Ma1 \ unit_disc" "Ma2 \ unit_disc" "?P' Ma1 Ma2 (?M u) (?M x1) (?M x2)" using **[rule_format, OF MM] by blast hence MM': "\a'\unit_disc. poincare_in_angle a' Ma1 (?M u) Ma2 \ poincare_ray_meets_line (?M u) a' (?M x1) (?M x2)" by auto obtain a1 a2 where a: "a1 \ unit_disc" "a2 \ unit_disc" "?M a1 = Ma1" "?M a2 = Ma2" using uf by (metis \Ma1 \ unit_disc\ \Ma2 \ unit_disc\ image_iff unit_disc_fix_iff) have "\a'\unit_disc. poincare_in_angle a' a1 u a2 \ poincare_ray_meets_line u a' x1 x2" proof safe fix a' assume "a' \ unit_disc" "poincare_in_angle a' a1 u a2" thus "poincare_ray_meets_line u a' x1 x2" using MM(1-2) MM'[rule_format, of "?M a'"] * uf a xx by (meson unit_disc_fix_discI unit_disc_fix_preserves_poincare_in_angle unit_disc_fix_preserves_poincare_ray_meets_line) qed hence "?P' a1 a2 u x1 x2" using MM * uf xx a by auto thus ?thesis using \a1 \ unit_disc\ \a2 \ unit_disc\ by blast qed qed next show "?Q 0\<^sub>h" proof safe fix x1 x2 assume "x1 \ unit_disc" "x2 \ unit_disc" assume "\ poincare_collinear {0\<^sub>h, x1, x2}" show "?P 0\<^sub>h x1 x2" proof- let ?lx = "poincare_line x1 x2" have "x1 \ x2" using \x1 \ unit_disc\ \x2 \ unit_disc\\\ poincare_collinear {0\<^sub>h, x1, x2}\ using poincare_collinear3_between by auto have lx: "is_poincare_line ?lx" using is_poincare_line_poincare_line[OF \x1 \ x2\] by simp obtain i1 i2 where "ideal_points ?lx = {i1, i2}" by (meson \x1 \ x2\ is_poincare_line_poincare_line obtain_ideal_points) let ?li = "poincare_line i1 i2" let ?i1 = "to_complex i1" let ?i2 = "to_complex i2" have "i1 \ unit_circle_set" "i2 \ unit_circle_set" using lx \ideal_points ?lx = {i1, i2}\ unfolding unit_circle_set_def by (metis ideal_points_on_unit_circle insertI1, metis ideal_points_on_unit_circle insertI1 insertI2) have "i1 \ i2" using \ideal_points ?lx = {i1, i2}\ \x1 \ unit_disc\ \x1 \ x2\ \x2 \ unit_disc\ ideal_points_different(1) by blast let ?a1 = "of_complex (?i1 / 2)" let ?a2 = "of_complex (?i2 / 2)" let ?la = "poincare_line ?a1 ?a2" have "?a1 \ unit_disc" "?a2 \ unit_disc" using \i1 \ unit_circle_set\ \i2 \ unit_circle_set\ unfolding unit_circle_set_def unit_disc_def disc_def circline_set_def by auto (transfer, transfer, case_tac i1, case_tac i2, simp add: vec_cnj_def)+ have "?a1 \ 0\<^sub>h" "?a2 \ 0\<^sub>h" using \i1 \ unit_circle_set\ \i2 \ unit_circle_set\ unfolding unit_circle_set_def by auto have "?a1 \ ?a2" using \i1 \ i2\ by (metis \i1 \ unit_circle_set\ \i2 \ unit_circle_set\ circline_set_def divide_cancel_right inversion_infty inversion_unit_circle mem_Collect_eq of_complex_to_complex of_complex_zero to_complex_of_complex unit_circle_set_def zero_neq_numeral) have "poincare_collinear {0\<^sub>h, ?a1, i1}" unfolding poincare_collinear_def using \?a1 \ 0\<^sub>h\[symmetric] is_poincare_line_poincare_line[of "0\<^sub>h" ?a1] unfolding circline_set_def apply (rule_tac x="poincare_line 0\<^sub>h ?a1" in exI, auto) apply (transfer, transfer, auto simp add: vec_cnj_def) done have "poincare_collinear {0\<^sub>h, ?a2, i2}" unfolding poincare_collinear_def using \?a2 \ 0\<^sub>h\[symmetric] is_poincare_line_poincare_line[of "0\<^sub>h" ?a2] unfolding circline_set_def apply (rule_tac x="poincare_line 0\<^sub>h ?a2" in exI, auto) apply (transfer, transfer, auto simp add: vec_cnj_def) done have "\ poincare_ray_meets_line 0\<^sub>h ?a1 x1 x2" using tangent_not_meet[of x1 x2 i1 ?a1] using \x1 \ unit_disc\ \x2 \ unit_disc\ \?a1 \ unit_disc\ \x1 \ x2\ \\ poincare_collinear {0\<^sub>h, x1, x2}\ using \ideal_points ?lx = {i1, i2}\ \?a1 \ 0\<^sub>h\ \poincare_collinear {0\<^sub>h, ?a1, i1}\ by simp moreover have "\ poincare_ray_meets_line 0\<^sub>h ?a2 x1 x2" using tangent_not_meet[of x1 x2 i2 ?a2] using \x1 \ unit_disc\ \x2 \ unit_disc\ \?a2 \ unit_disc\ \x1 \ x2\ \\ poincare_collinear {0\<^sub>h, x1, x2}\ using \ideal_points ?lx = {i1, i2}\ \?a2 \ 0\<^sub>h\ \poincare_collinear {0\<^sub>h, ?a2, i2}\ by simp moreover have "\a' \ unit_disc. poincare_in_angle a' ?a1 0\<^sub>h ?a2 \ poincare_ray_meets_line 0\<^sub>h a' x1 x2" unfolding poincare_in_angle_def proof safe fix a' a assume *: "a' \ unit_disc" "a \ unit_disc" "poincare_on_ray a' 0\<^sub>h a" "a' \ 0\<^sub>h" "poincare_between ?a1 a ?a2" "a \ ?a1" "a \ ?a2" show "poincare_ray_meets_line 0\<^sub>h a' x1 x2" proof- have "\ a' a1 a2 x1 x2 i1 i2. a' \ unit_disc \ x1 \ unit_disc \ x2 \ unit_disc \ x1 \ x2 \ \ poincare_collinear {0\<^sub>h, x1, x2} \ ideal_points (poincare_line x1 x2) = {i1, i2} \ a1 = of_complex (to_complex i1 / 2) \ a2 = of_complex (to_complex i2 / 2) \ i1 \ i2 \ a1 \ a2 \ poincare_collinear {0\<^sub>h, a1, i1} \ poincare_collinear {0\<^sub>h, a2, i2} \ a1 \ unit_disc \ a2 \ unit_disc \ i1 \ unit_circle_set \ i2 \ unit_circle_set \ poincare_on_ray a' 0\<^sub>h a \ a' \ 0\<^sub>h \ poincare_between a1 a a2 \ a \ a1 \ a \ a2 \ poincare_ray_meets_line 0\<^sub>h a' x1 x2" (is "\ a' a1 a2 x1 x2 i1 i2. ?R 0\<^sub>h a' a1 a2 x1 x2 i1 i2 a") proof (rule wlog_rotation_to_positive_x_axis[OF \a \ unit_disc\]) let ?R' = "\ a zero. \ a' a1 a2 x1 x2 i1 i2. ?R zero a' a1 a2 x1 x2 i1 i2 a" fix xa assume xa: "is_real xa" "0 < Re xa" "Re xa < 1" let ?a = "of_complex xa" show "?R' ?a 0\<^sub>h" proof safe fix a' a1 a2 x1 x2 i1 i2 let ?i1 = "to_complex i1" and ?i2 = "to_complex i2" let ?a1 = "of_complex (?i1 / 2)" and ?a2 = "of_complex (?i2 / 2)" let ?la = "poincare_line ?a1 ?a2" and ?lx = "poincare_line x1 x2" and ?li = "poincare_line i1 i2" assume "a' \ unit_disc" "x1 \ unit_disc" "x2 \ unit_disc" "x1 \ x2" assume "\ poincare_collinear {0\<^sub>h, x1, x2}" "ideal_points ?lx = {i1, i2}" assume "poincare_on_ray a' 0\<^sub>h ?a" "a' \ 0\<^sub>h" assume "poincare_between ?a1 ?a ?a2" "?a \ ?a1" "?a \ ?a2" assume "i1 \ i2" "?a1 \ ?a2" "poincare_collinear {0\<^sub>h, ?a1, i1}" "poincare_collinear {0\<^sub>h, ?a2, i2}" assume "?a1 \ unit_disc" "?a2 \ unit_disc" assume "i1 \ unit_circle_set" "i2 \ unit_circle_set" show "poincare_ray_meets_line 0\<^sub>h a' x1 x2" proof- have "?lx = ?li" using \ideal_points ?lx = {i1, i2}\ \x1 \ x2\ ideal_points_line_unique by auto have lx: "is_poincare_line ?lx" using is_poincare_line_poincare_line[OF \x1 \ x2\] by simp have "x1 \ circline_set ?lx" "x2 \ circline_set ?lx" using lx \x1 \ x2\ by auto have "?lx \ x_axis" using \\ poincare_collinear {0\<^sub>h, x1, x2}\ \x1 \ circline_set ?lx\ \x2 \ circline_set ?lx\ lx unfolding poincare_collinear_def by auto have "0\<^sub>h \ circline_set ?lx" using \\ poincare_collinear {0\<^sub>h, x1, x2}\ lx \x1 \ circline_set ?lx\ \x2 \ circline_set ?lx\ unfolding poincare_collinear_def by auto have "xa \ 0" "?a \ 0\<^sub>h" using xa by auto hence "0\<^sub>h \ ?a" by metis have "?a \ positive_x_axis" using xa unfolding positive_x_axis_def by simp have "?a \ unit_disc" using xa by (auto simp add: cmod_eq_Re) have "?a \ circline_set ?la" using \poincare_between ?a1 ?a ?a2\ using \?a1 \ ?a2\ \?a \ unit_disc\ \?a1 \ unit_disc\ \?a2 \ unit_disc\ poincare_between_poincare_line_uzv by blast have "?a1 \ circline_set ?la" "?a2 \ circline_set ?la" by (auto simp add: \?a1 \ ?a2\) have la: "is_poincare_line ?la" using is_poincare_line_poincare_line[OF \?a1 \ ?a2\] by simp have inv: "inversion i1 = i1" "inversion i2 = i2" using \i1 \ unit_circle_set\ \i2 \ unit_circle_set\ by (auto simp add: circline_set_def unit_circle_set_def) have "i1 \ \\<^sub>h" "i2 \ \\<^sub>h" using inv by auto have "?a1 \ circline_set x_axis \ ?a2 \ circline_set x_axis" proof (rule ccontr) assume "\ ?thesis" hence "?a1 \ circline_set x_axis \ ?a2 \ circline_set x_axis" by auto hence "?la = x_axis" proof assume "?a1 \ circline_set x_axis" hence "{?a, ?a1} \ circline_set ?la \ circline_set x_axis" using \?a \ circline_set ?la\ \?a1 \ circline_set ?la\\?a \ positive_x_axis\ using circline_set_x_axis_I xa(1) by blast thus "?la = x_axis" using unique_is_poincare_line[of ?a ?a1 ?la x_axis] using \?a1 \ unit_disc\ \?a \ unit_disc\ la \?a \ ?a1\ by auto next assume "?a2 \ circline_set x_axis" hence "{?a, ?a2} \ circline_set ?la \ circline_set x_axis" using \?a \ circline_set ?la\ \?a2 \ circline_set ?la\ \?a \ positive_x_axis\ using circline_set_x_axis_I xa(1) by blast thus "?la = x_axis" using unique_is_poincare_line[of ?a ?a2 ?la x_axis] using \?a2 \ unit_disc\ \?a \ unit_disc\ la \?a \ ?a2\ by auto qed hence "i1 \ circline_set x_axis \ i2 \ circline_set x_axis" using \?a1 \ circline_set ?la\ \?a2 \ circline_set ?la\ by (metis \i1 \ \\<^sub>h\ \i2 \ \\<^sub>h\ \of_complex (to_complex i1 / 2) \ unit_disc\ \of_complex (to_complex i2 / 2) \ unit_disc\ \poincare_collinear {0\<^sub>h, of_complex (to_complex i1 / 2), i1}\ \poincare_collinear {0\<^sub>h, of_complex (to_complex i2 / 2), i2}\ divide_eq_0_iff inf_not_of_complex inv(1) inv(2) inversion_noteq_unit_disc of_complex_to_complex of_complex_zero_iff poincare_collinear3_poincare_lines_equal_general poincare_line_0_real_is_x_axis poincare_line_circline_set(2) zero_in_unit_disc zero_neq_numeral) thus False using \?lx \ x_axis\ unique_is_poincare_line_general[of i1 i2 ?li x_axis] \i1 \ i2\ inv \?lx = ?li\ by auto qed hence "?la \ x_axis" using \?a1 \ ?a2\ poincare_line_circline_set(1) by fastforce have "intersects_x_axis_positive ?la" using intersects_x_axis_positive_iff[of ?la] \?la \ x_axis\ \?a \ circline_set ?la\ la using \?a \ unit_disc\ \?a \ positive_x_axis\ by auto have "intersects_x_axis ?lx" proof- have "arg (to_complex ?a1) * arg (to_complex ?a2) < 0" using \poincare_between ?a1 ?a ?a2\ \?a1 \ unit_disc\ \?a2 \ unit_disc\ using poincare_between_x_axis_intersection[of ?a1 ?a2 "of_complex xa"] using \?a1 \ ?a2\ \?a \ unit_disc\ \?a1 \ circline_set x_axis \ ?a2 \ circline_set x_axis\ \?a \ positive_x_axis\ using \?a \ circline_set ?la\ unfolding positive_x_axis_def by simp moreover have "\ x y x' y' :: real. \sgn x' = sgn x; sgn y' = sgn y\ \ x*y < 0 \ x'*y' < 0" by (metis sgn_less sgn_mult) ultimately have "Im (to_complex ?a1) * Im (to_complex ?a2) < 0" using arg_Im_sgn[of "to_complex ?a1"] arg_Im_sgn[of "to_complex ?a2"] using \?a1 \ unit_disc\ \?a2 \ unit_disc\ \?a1 \ circline_set x_axis \ ?a2 \ circline_set x_axis\ using inf_or_of_complex[of ?a1] inf_or_of_complex[of ?a2] circline_set_x_axis by (metis circline_set_x_axis_I to_complex_of_complex) thus ?thesis using ideal_points_intersects_x_axis[of ?lx i1 i2] using \ideal_points ?lx = {i1, i2}\ lx \?lx \ x_axis\ by simp qed have "intersects_x_axis_positive ?lx" proof- have "cmod ?i1 = 1" "cmod ?i2 = 1" using \i1 \ unit_circle_set\ \i2 \ unit_circle_set\ unfolding unit_circle_set_def by auto let ?a1' = "?i1 / 2" and ?a2' = "?i2 / 2" let ?Aa1 = "\ * (?a1' * cnj ?a2' - ?a2' * cnj ?a1')" and ?Ba1 = "\ * (?a2' * cor ((cmod ?a1')\<^sup>2 + 1) - ?a1' * cor ((cmod ?a2')\<^sup>2 + 1))" have "?Aa1 \ 0 \ ?Ba1 \ 0" using \cmod (to_complex i1) = 1\ \cmod (to_complex i2) = 1\ \?a1 \ ?a2\ by (auto simp add: power_divide complex_mult_cnj_cmod) have "is_real ?Aa1" by simp have "?a1 \ inversion ?a2" using \?a1 \ unit_disc\ \?a2 \ unit_disc\ inversion_noteq_unit_disc by fastforce hence "Re ?Ba1 / Re ?Aa1 < -1" using \intersects_x_axis_positive ?la\ \?a1 \ ?a2\ using intersects_x_axis_positive_mk_circline[of ?Aa1 ?Ba1] \?Aa1 \ 0 \ ?Ba1 \ 0\ \is_real ?Aa1\ using poincare_line_non_homogenous[of ?a1 ?a2] by (simp add: Let_def) moreover let ?i1' = "to_complex i1" and ?i2' = "to_complex i2" let ?Ai1 = "\ * (?i1' * cnj ?i2' - ?i2' * cnj ?i1')" and ?Bi1 = "\ * (?i2' * cor ((cmod ?i1')\<^sup>2 + 1) - ?i1' * cor ((cmod ?i2')\<^sup>2 + 1))" have "?Ai1 \ 0 \ ?Bi1 \ 0" using \cmod (to_complex i1) = 1\ \cmod (to_complex i2) = 1\ \?a1 \ ?a2\ by (auto simp add: power_divide complex_mult_cnj_cmod) have "is_real ?Ai1" by simp have "sgn (Re ?Bi1 / Re ?Ai1) = sgn (Re ?Ba1 / Re ?Aa1)" proof- have "Re ?Bi1 / Re ?Ai1 = (Im ?i1 * 2 - Im ?i2 * 2) / (Im ?i2 * (Re ?i1 * 2) - Im ?i1 * (Re ?i2 * 2))" using \cmod ?i1 = 1\ \cmod ?i2 = 1\ by (auto simp add: complex_mult_cnj_cmod field_simps) also have "... = (Im ?i1 - Im ?i2) / (Im ?i2 * (Re ?i1) - Im ?i1 * (Re ?i2))" (is "... = ?expr") apply (subst left_diff_distrib[symmetric]) apply (subst semiring_normalization_rules(18))+ apply (subst left_diff_distrib[symmetric]) by (metis mult.commute mult_divide_mult_cancel_left_if zero_neq_numeral) finally have 1: "Re ?Bi1 / Re ?Ai1 = (Im ?i1 - Im ?i2) / (Im ?i2 * (Re ?i1) - Im ?i1 * (Re ?i2))" . have "Re ?Ba1 / Re ?Aa1 = (Im ?i1 * 20 - Im ?i2 * 20) / (Im ?i2 * (Re ?i1 * 16) - Im ?i1 * (Re ?i2 * 16))" using \cmod (to_complex i1) = 1\ \cmod (to_complex i2) = 1\ by (auto simp add: complex_mult_cnj_cmod field_simps) also have "... = (20 / 16) * ((Im ?i1 - Im ?i2) / (Im ?i2 * (Re ?i1) - Im ?i1 * (Re ?i2)))" apply (subst left_diff_distrib[symmetric])+ apply (subst semiring_normalization_rules(18))+ apply (subst left_diff_distrib[symmetric])+ by (metis (no_types, hide_lams) field_class.field_divide_inverse mult.commute times_divide_times_eq) finally have 2: "Re ?Ba1 / Re ?Aa1 = (5 / 4) * ((Im ?i1 - Im ?i2) / (Im ?i2 * (Re ?i1) - Im ?i1 * (Re ?i2)))" by simp have "?expr \ 0" using \Re ?Ba1 / Re ?Aa1 < -1\ apply (subst (asm) 2) by linarith thus ?thesis apply (subst 1, subst 2) apply (simp only: sgn_mult) by simp qed moreover have "i1 \ inversion i2" by (simp add: \i1 \ i2\ inv(2)) have "(Re ?Bi1 / Re ?Ai1)\<^sup>2 > 1" proof- have "?Ai1 = 0 \ (Re ?Bi1)\<^sup>2 > (Re ?Ai1)\<^sup>2" using \intersects_x_axis ?lx\ using \i1 \ i2\ \i1 \ \\<^sub>h\ \i2 \ \\<^sub>h\ \i1 \ inversion i2\ using intersects_x_axis_mk_circline[of ?Ai1 ?Bi1] \?Ai1 \ 0 \ ?Bi1 \ 0\ \is_real ?Ai1\ using poincare_line_non_homogenous[of i1 i2] \?lx = ?li\ by metis moreover have "?Ai1 \ 0" proof (rule ccontr) assume "\ ?thesis" hence "0\<^sub>h \ circline_set ?li" unfolding circline_set_def apply simp apply (transfer, transfer, case_tac i1, case_tac i2) by (auto simp add: vec_cnj_def field_simps) thus False using \0\<^sub>h \ circline_set ?lx\ \?lx = ?li\ by simp qed ultimately have "(Re ?Bi1)\<^sup>2 > (Re ?Ai1)\<^sup>2" by auto moreover have "Re ?Ai1 \ 0" using \is_real ?Ai1\ \?Ai1 \ 0\ by (simp add: complex_eq_iff) ultimately show ?thesis by (simp add: power_divide) qed moreover { fix x1 x2 :: real assume "sgn x1 = sgn x2" "x1 < -1" "x2\<^sup>2 > 1" hence "x2 < -1" by (smt one_power2 real_sqrt_abs real_sqrt_less_iff sgn_neg sgn_pos) } ultimately have "Re ?Bi1 / Re ?Ai1 < -1" by metis thus ?thesis using \i1 \ i2\ \i1 \ \\<^sub>h\ \i2 \ \\<^sub>h\ \i1 \ inversion i2\ using intersects_x_axis_positive_mk_circline[of ?Ai1 ?Bi1] \?Ai1 \ 0 \ ?Bi1 \ 0\ \is_real ?Ai1\ using poincare_line_non_homogenous[of i1 i2] \?lx = ?li\ by (simp add: Let_def) qed then obtain x where x: "x \ unit_disc" "x \ circline_set ?lx \ positive_x_axis" using intersects_x_axis_positive_iff[OF lx \?lx \ x_axis\] by auto have "poincare_on_ray x 0\<^sub>h a' \ poincare_collinear {x1, x2, x}" proof show "poincare_collinear {x1, x2, x}" using x lx \x1 \ circline_set ?lx\ \x2 \ circline_set ?lx\ unfolding poincare_collinear_def by auto next show "poincare_on_ray x 0\<^sub>h a'" unfolding poincare_on_ray_def proof- have "a' \ circline_set x_axis" using \poincare_on_ray a' 0\<^sub>h ?a\ xa \0\<^sub>h \ ?a\ \xa \ 0\ \a' \ unit_disc\ unfolding poincare_on_ray_def using poincare_line_0_real_is_x_axis[of "of_complex xa"] using poincare_between_poincare_line_uvz[of "0\<^sub>h" "of_complex xa" a'] using poincare_between_poincare_line_uzv[of "0\<^sub>h" "of_complex xa" a'] by (auto simp add: cmod_eq_Re) then obtain xa' where xa': "a' = of_complex xa'" "is_real xa'" using \a' \ unit_disc\ using circline_set_def on_circline_x_axis by auto hence "-1 < Re xa'" "Re xa' < 1" "xa' \ 0" using \a' \ unit_disc\ \a' \ 0\<^sub>h\ by (auto simp add: cmod_eq_Re) hence "Re xa' > 0" "Re xa' < 1" "is_real xa'" using \poincare_on_ray a' 0\<^sub>h (of_complex xa)\ using poincare_between_x_axis_0uv[of "Re xa'" "Re xa"] using poincare_between_x_axis_0uv[of "Re xa" "Re xa'"] using circline_set_positive_x_axis_I[of "Re xa'"] using xa xa' complex_of_real_Re unfolding poincare_on_ray_def by (smt of_real_0, linarith, blast) moreover obtain xx where "is_real xx" "Re xx > 0" "Re xx < 1" "x = of_complex xx" using x unfolding positive_x_axis_def using circline_set_def cmod_eq_Re on_circline_x_axis by auto ultimately show "poincare_between 0\<^sub>h x a' \ poincare_between 0\<^sub>h a' x" using \a' = of_complex xa'\ by (smt \a' \ unit_disc\ arg_0_iff poincare_between_0uv poincare_between_def to_complex_of_complex x(1)) qed qed thus ?thesis using \x \ unit_disc\ unfolding poincare_ray_meets_line_def poincare_on_line_def by (metis insert_commute) qed qed next show "a \ 0\<^sub>h" proof (rule ccontr) assume "\ ?thesis" then obtain k where "k<0" "to_complex ?a1 = cor k * to_complex ?a2" using poincare_between_u0v[OF \?a1 \ unit_disc\ \?a2 \ unit_disc\ \?a1 \ 0\<^sub>h\ \?a2 \ 0\<^sub>h\] using \poincare_between ?a1 a ?a2\ by auto hence "to_complex i1 = cor k * to_complex i2" "k < 0" by auto hence "0\<^sub>h \ circline_set (poincare_line x1 x2)" using ideal_points_proportional[of "poincare_line x1 x2" i1 i2 k] \ideal_points (poincare_line x1 x2) = {i1, i2}\ using is_poincare_line_poincare_line[OF \x1 \ x2\] by simp thus False using \\ poincare_collinear {0\<^sub>h, x1, x2}\ using is_poincare_line_poincare_line[OF \x1 \ x2\] unfolding poincare_collinear_def by (meson \x1 \ x2\ empty_subsetI insert_subset poincare_line_circline_set(1) poincare_line_circline_set(2)) qed next fix \ u let ?R' = "\ a zero. \ a' a1 a2 x1 x2 i1 i2. ?R zero a' a1 a2 x1 x2 i1 i2 a" let ?M = "moebius_pt (moebius_rotation \)" assume *: "u \ unit_disc" "u \ 0\<^sub>h" and **: "?R' (?M u) 0\<^sub>h" have uf: "unit_disc_fix (moebius_rotation \)" by simp have "?M 0\<^sub>h = 0\<^sub>h" by auto hence **: "?R' (?M u) (?M 0\<^sub>h)" using ** by simp show "?R' u 0\<^sub>h" proof (rule allI)+ fix a' a1 a2 x1 x2 i1 i2 have i1: "i1 \ unit_circle_set \ moebius_pt (moebius_rotation \) (of_complex (to_complex i1 / 2)) = of_complex (to_complex (moebius_pt (moebius_rotation \) i1) / 2)" using unit_circle_set_def by force have i2: "i2 \ unit_circle_set \ moebius_pt (moebius_rotation \) (of_complex (to_complex i2 / 2)) = of_complex (to_complex (moebius_pt (moebius_rotation \) i2) / 2)" using unit_circle_set_def by force show "?R 0\<^sub>h a' a1 a2 x1 x2 i1 i2 u" using **[rule_format, of "?M a'" "?M x1" "?M x2" "?M i1" "?M i2" "?M a1" "?M a2"] uf * apply (auto simp del: moebius_pt_moebius_rotation_zero moebius_pt_moebius_rotation) using i1 i2 by simp qed qed thus ?thesis using \a' \ unit_disc\ \x1 \ unit_disc\ \x2 \ unit_disc\ \x1 \ x2\ using \\ poincare_collinear {0\<^sub>h, x1, x2}\ \ideal_points ?lx = {i1, i2}\ \i1 \ i2\ using \?a1 \ ?a2\ \poincare_collinear {0\<^sub>h, ?a1, i1}\ \poincare_collinear {0\<^sub>h, ?a2, i2}\ using \?a1 \ unit_disc\ \?a2 \ unit_disc\ \i1 \ unit_circle_set\ \i2 \ unit_circle_set\ using \poincare_on_ray a' 0\<^sub>h a\ \a' \ 0\<^sub>h\ \poincare_between ?a1 a ?a2\ \a \ ?a1\ \a \ ?a2\ by blast qed qed moreover have "\ poincare_on_line 0\<^sub>h ?a1 ?a2" proof assume *: "poincare_on_line 0\<^sub>h ?a1 ?a2" hence "poincare_collinear {0\<^sub>h, ?a1, ?a2}" unfolding poincare_on_line_def by simp hence "poincare_line 0\<^sub>h ?a1 = poincare_line 0\<^sub>h ?a2" using poincare_collinear3_poincare_lines_equal_general[of "0\<^sub>h" ?a1 ?a2] using \?a1 \ unit_disc\ \?a1 \ 0\<^sub>h\ \?a2 \ unit_disc\ \?a2 \ 0\<^sub>h\ by (metis inversion_noteq_unit_disc zero_in_unit_disc) have "i1 \ circline_set (poincare_line 0\<^sub>h ?a1)" using \poincare_collinear {0\<^sub>h, ?a1, i1}\ using poincare_collinear3_poincare_line_general[of i1 "0\<^sub>h" ?a1] using \?a1 \ unit_disc\ \?a1 \ 0\<^sub>h\ by (metis insert_commute inversion_noteq_unit_disc zero_in_unit_disc) moreover have "i2 \ circline_set (poincare_line 0\<^sub>h ?a1)" using \poincare_collinear {0\<^sub>h, ?a2, i2}\ using poincare_collinear3_poincare_line_general[of i2 "0\<^sub>h" ?a2] using \?a2 \ unit_disc\ \?a2 \ 0\<^sub>h\ \poincare_line 0\<^sub>h ?a1 = poincare_line 0\<^sub>h ?a2\ by (metis insert_commute inversion_noteq_unit_disc zero_in_unit_disc) ultimately have "poincare_collinear {0\<^sub>h, i1, i2}" using \?a1 \ unit_disc\ \?a1 \ 0\<^sub>h\ \poincare_collinear {0\<^sub>h, ?a1, i1}\ by (smt insert_subset poincare_collinear_def unique_poincare_line zero_in_unit_disc) hence "0\<^sub>h \ circline_set (poincare_line i1 i2)" using poincare_collinear3_poincare_line_general[of "0\<^sub>h" i1 i2] using \i1 \ i2\ \i2 \ unit_circle_set\ unit_circle_set_def by force moreover have "?lx = ?li" using \ideal_points ?lx = {i1, i2}\ \x1 \ x2\ ideal_points_line_unique by auto ultimately show False using \\ poincare_collinear {0\<^sub>h, x1, x2}\ using \x1 \ x2\ poincare_line_poincare_collinear3_general by auto qed ultimately show ?thesis using \?a1 \ unit_disc\ \?a2 \ unit_disc\ by blast qed qed qed thus ?thesis using \x1 \ unit_disc\ \x2 \ unit_disc\ \\ poincare_collinear {a, x1, x2}\ by blast qed subsection\Interpretation of locales\ global_interpretation PoincareTarskiAbsolute: TarskiAbsolute where cong = p_congruent and betw = p_between defines p_on_line = PoincareTarskiAbsolute.on_line and p_on_ray = PoincareTarskiAbsolute.on_ray and p_in_angle = PoincareTarskiAbsolute.in_angle and p_ray_meets_line = PoincareTarskiAbsolute.ray_meets_line proof- show "TarskiAbsolute p_congruent p_between" proof text\ 1. Reflexivity of congruence \ fix x y show "p_congruent x y y x" unfolding p_congruent_def by transfer (simp add: poincare_distance_sym) next text\ 2. Transitivity of congruence \ fix x y z u v w show "p_congruent x y z u \ p_congruent x y v w \ p_congruent z u v w" by (transfer, simp) next text\ 3. Identity of congruence \ fix x y z show "p_congruent x y z z \ x = y" unfolding p_congruent_def by transfer (simp add: poincare_distance_eq_0_iff) next text\ 4. Segment construction \ fix x y a b show "\ z. p_between x y z \ p_congruent y z a b" using segment_construction unfolding p_congruent_def by transfer (simp, blast) next text\ 5. Five segment \ fix x y z x' y' z' u u' show "x \ y \ p_between x y z \ p_between x' y' z' \ p_congruent x y x' y' \ p_congruent y z y' z' \ p_congruent x u x' u' \ p_congruent y u y' u' \ p_congruent z u z' u'" unfolding p_congruent_def apply transfer using five_segment_axiom by meson next text\ 6. Identity of betweeness \ fix x y show "p_between x y x \ x = y" by transfer (simp add: poincare_between_sum_distances poincare_distance_eq_0_iff poincare_distance_sym) next text\ 7. Pasch \ fix x y z u v show "p_between x u z \ p_between y v z \ (\ a. p_between u a y \ p_between x a v)" apply transfer using Pasch by blast next text\ 8. Lower dimension \ show "\ a. \ b. \ c. \ p_between a b c \ \ p_between b c a \ \ p_between c a b" apply (transfer) using lower_dimension_axiom by simp next text\ 9. Upper dimension \ fix x y z u v show "p_congruent x u x v \ p_congruent y u y v \ p_congruent z u z v \ u \ v \ p_between x y z \ p_between y z x \ p_between z x y" unfolding p_congruent_def by (transfer, simp add: upper_dimension_axiom) qed qed interpretation PoincareTarskiHyperbolic: TarskiHyperbolic where cong = p_congruent and betw = p_between proof text\ 10. Euclid negation \ show "\ a b c d t. p_between a d t \ p_between b d c \ a \ d \ (\ x y. p_between a b x \ p_between a c y \ \ p_between x t y)" using negated_euclidean_axiom by transfer (auto, blast) next fix a x1 x2 assume "\ TarskiAbsolute.on_line p_between a x1 x2" hence "\ p_on_line a x1 x2" using TarskiAbsolute.on_line_def[OF PoincareTarskiAbsolute.TarskiAbsolute_axioms] using PoincareTarskiAbsolute.on_line_def by simp text\ 11. Limiting parallels \ thus "\a1 a2. \ TarskiAbsolute.on_line p_between a a1 a2 \ \ TarskiAbsolute.ray_meets_line p_between a a1 x1 x2 \ \ TarskiAbsolute.ray_meets_line p_between a a2 x1 x2 \ (\a'. TarskiAbsolute.in_angle p_between a' a1 a a2 \ TarskiAbsolute.ray_meets_line p_between a a' x1 x2)" unfolding TarskiAbsolute.in_angle_def[OF PoincareTarskiAbsolute.TarskiAbsolute_axioms] unfolding TarskiAbsolute.on_ray_def[OF PoincareTarskiAbsolute.TarskiAbsolute_axioms] unfolding TarskiAbsolute.ray_meets_line_def[OF PoincareTarskiAbsolute.TarskiAbsolute_axioms] unfolding TarskiAbsolute.on_ray_def[OF PoincareTarskiAbsolute.TarskiAbsolute_axioms] unfolding TarskiAbsolute.on_line_def[OF PoincareTarskiAbsolute.TarskiAbsolute_axioms] unfolding PoincareTarskiAbsolute.on_line_def apply transfer proof- fix a x1 x2 assume *: "a \ unit_disc" "x1 \ unit_disc" "x2 \ unit_disc" "\ (poincare_between a x1 x2 \ poincare_between x1 a x2 \ poincare_between x1 x2 a)" hence "\ poincare_on_line a x1 x2" using poincare_collinear3_iff[of a x1 x2] using poincare_between_rev poincare_on_line_def by blast hence "\a1\unit_disc. \a2\unit_disc. \ poincare_on_line a a1 a2 \ \ poincare_ray_meets_line a a1 x1 x2 \ \ poincare_ray_meets_line a a2 x1 x2 \ (\a'\unit_disc. poincare_in_angle a' a1 a a2 \ poincare_ray_meets_line a a' x1 x2)" using limiting_parallels[of a x1 x2] * by blast then obtain a1 a2 where **: "a1\unit_disc" "a2\unit_disc" "\ poincare_on_line a a1 a2" "\ poincare_ray_meets_line a a2 x1 x2" "\ poincare_ray_meets_line a a1 x1 x2" "\a'\unit_disc. poincare_in_angle a' a1 a a2 \ poincare_ray_meets_line a a' x1 x2" by blast have "\ (\x\{z. z \ unit_disc}. (poincare_between a x a1 \ poincare_between a a1 x) \ (poincare_between x x1 x2 \ poincare_between x1 x x2 \ poincare_between x1 x2 x))" using \\ poincare_ray_meets_line a a1 x1 x2\ unfolding poincare_on_line_def poincare_ray_meets_line_def poincare_on_ray_def using poincare_collinear3_iff[of _ x1 x2] poincare_between_rev *(2, 3) by auto moreover have "\ (\x\{z. z \ unit_disc}. (poincare_between a x a2 \ poincare_between a a2 x) \ (poincare_between x x1 x2 \ poincare_between x1 x x2 \ poincare_between x1 x2 x))" using \\ poincare_ray_meets_line a a2 x1 x2\ unfolding poincare_on_line_def poincare_ray_meets_line_def poincare_on_ray_def using poincare_collinear3_iff[of _ x1 x2] poincare_between_rev *(2, 3) by auto moreover have "\ (poincare_between a a1 a2 \ poincare_between a1 a a2 \ poincare_between a1 a2 a)" using \\ poincare_on_line a a1 a2\ poincare_collinear3_iff[of a a1 a2] using *(1) **(1-2) unfolding poincare_on_line_def by simp moreover have "(\a'\{z. z \ unit_disc}. a \ a1 \ a \ a2 \ a' \ a \ (\x\{z. z \ unit_disc}. poincare_between a1 x a2 \ x \ a1 \ x \ a2 \ (poincare_between a a' x \ poincare_between a x a')) \ (\x\{z. z \ unit_disc}. (poincare_between a x a' \ poincare_between a a' x) \ (poincare_between x x1 x2 \ poincare_between x1 x x2 \ poincare_between x1 x2 x)))" using **(6) unfolding poincare_on_line_def poincare_in_angle_def poincare_ray_meets_line_def poincare_on_ray_def using poincare_collinear3_iff[of _ x1 x2] poincare_between_rev *(2, 3) by auto ultimately show "\a1\{z. z \ unit_disc}. \a2\{z. z \ unit_disc}. \ (poincare_between a a1 a2 \ poincare_between a1 a a2 \ poincare_between a1 a2 a) \ \ (\x\{z. z \ unit_disc}. (poincare_between a x a1 \ poincare_between a a1 x) \ (poincare_between x x1 x2 \ poincare_between x1 x x2 \ poincare_between x1 x2 x)) \ \ (\x\{z. z \ unit_disc}. (poincare_between a x a2 \ poincare_between a a2 x) \ (poincare_between x x1 x2 \ poincare_between x1 x x2 \ poincare_between x1 x2 x)) \ (\a'\{z. z \ unit_disc}. a \ a1 \ a \ a2 \ a' \ a \ (\x\{z. z \ unit_disc}. poincare_between a1 x a2 \ x \ a1 \ x \ a2 \ (poincare_between a a' x \ poincare_between a x a')) \ (\x\{z. z \ unit_disc}. (poincare_between a x a' \ poincare_between a a' x) \ (poincare_between x x1 x2 \ poincare_between x1 x x2 \ poincare_between x1 x2 x)))" using **(1, 2) by auto qed qed interpretation PoincareElementaryTarskiHyperbolic: ElementaryTarskiHyperbolic p_congruent p_between proof text\ 12. Continuity \ fix \ \ assume "\ a. \ x. \ y. \ x \ \ y \ p_between a x y" thus "\ b. \ x. \ y. \ x \ \ y \ p_between x b y" apply transfer using continuity by auto qed end diff --git a/thys/Stirling_Formula/Gamma_Asymptotics.thy b/thys/Stirling_Formula/Gamma_Asymptotics.thy --- a/thys/Stirling_Formula/Gamma_Asymptotics.thy +++ b/thys/Stirling_Formula/Gamma_Asymptotics.thy @@ -1,2105 +1,2105 @@ (* File: Gamma_Asymptotics.thy Author: Manuel Eberl The complete asymptotics of the real and complex logarithmic Gamma functions. Also of the real Polygamma functions (could be extended to the complex ones fairly easily if needed). *) section \Complete asymptotics of the logarithmic Gamma function\ theory Gamma_Asymptotics imports "HOL-Complex_Analysis.Complex_Analysis" "HOL-Real_Asymp.Real_Asymp" Bernoulli.Bernoulli_FPS Bernoulli.Periodic_Bernpoly Stirling_Formula begin subsection \Auxiliary Facts\ (* TODO Move *) lemma arg_of_real [simp]: - "x > 0 \ arg (complex_of_real x) = 0" - "x < 0 \ arg (complex_of_real x) = pi" - by (rule arg_unique; simp add: complex_sgn_def scaleR_conv_of_real)+ + "x > 0 \ Arg (complex_of_real x) = 0" + "x < 0 \ Arg (complex_of_real x) = pi" + by (auto simp add: Arg_of_real) lemma arg_conv_arctan: assumes "Re z > 0" - shows "arg z = arctan (Im z / Re z)" + shows "Arg z = arctan (Im z / Re z)" proof (rule arg_unique) show "sgn z = cis (arctan (Im z / Re z))" proof (rule complex_eqI) have "Re (cis (arctan (Im z / Re z))) = 1 / sqrt (1 + (Im z)\<^sup>2 / (Re z)\<^sup>2)" by (simp add: cos_arctan power_divide) also have "1 + Im z ^ 2 / Re z ^ 2 = norm z ^ 2 / Re z ^ 2" using assms by (simp add: cmod_def field_simps) also have "1 / sqrt \ = Re z / norm z" using assms by (simp add: real_sqrt_divide) finally show "Re (sgn z) = Re (cis (arctan (Im z / Re z)))" by simp next have "Im (cis (arctan (Im z / Re z))) = Im z / (Re z * sqrt (1 + (Im z)\<^sup>2 / (Re z)\<^sup>2))" by (simp add: sin_arctan field_simps) also have "1 + Im z ^ 2 / Re z ^ 2 = norm z ^ 2 / Re z ^ 2" using assms by (simp add: cmod_def field_simps) also have "Im z / (Re z * sqrt \) = Im z / norm z" using assms by (simp add: real_sqrt_divide) finally show "Im (sgn z) = Im (cis (arctan (Im z / Re z)))" by simp qed next show "arctan (Im z / Re z) > -pi" by (rule le_less_trans[OF _ arctan_lbound]) auto next have "arctan (Im z / Re z) < pi / 2" by (rule arctan_ubound) also have "\ \ pi" by simp finally show "arctan (Im z / Re z) \ pi" by simp qed lemma mult_indicator_cong: fixes f g :: "_ \ 'a :: semiring_1" shows "(\x. x \ A \ f x = g x) \ indicator A x * f x = indicator A x * g x" by (auto simp: indicator_def) lemma has_absolute_integral_change_of_variables_1': fixes f :: "real \ real" and g :: "real \ real" assumes S: "S \ sets lebesgue" and der_g: "\x. x \ S \ (g has_field_derivative g' x) (at x within S)" and inj: "inj_on g S" shows "(\x. \g' x\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \g' x\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" proof - have "(\x. \g' x\ *\<^sub>R vec (f(g x)) :: real ^ 1) absolutely_integrable_on S \ integral S (\x. \g' x\ *\<^sub>R vec (f(g x))) = (vec b :: real ^ 1) \ (\x. vec (f x) :: real ^ 1) absolutely_integrable_on (g ` S) \ integral (g ` S) (\x. vec (f x)) = (vec b :: real ^ 1)" using assms unfolding has_field_derivative_iff_has_vector_derivative by (intro has_absolute_integral_change_of_variables_1 assms) auto thus ?thesis by (simp add: absolutely_integrable_on_1_iff integral_on_1_eq) qed corollary Ln_times_of_nat: "\r > 0; z \ 0\ \ Ln(of_nat r * z :: complex) = ln (of_nat r) + Ln(z)" using Ln_times_of_real[of "of_nat r" z] by simp lemma tendsto_of_real_0_I: "(f \ 0) G \ ((\x. (of_real (f x))) \ (0 ::'a::real_normed_div_algebra)) G" by (subst (asm) tendsto_of_real_iff [symmetric]) simp lemma negligible_atLeastAtMostI: "b \ a \ negligible {a..(b::real)}" by (cases "b < a") auto lemma vector_derivative_cong_eq: assumes "eventually (\x. x \ A \ f x = g x) (nhds x)" "x = y" "A = B" "x \ A" shows "vector_derivative f (at x within A) = vector_derivative g (at y within B)" proof - from eventually_nhds_x_imp_x[OF assms(1)] assms(4) have "f x = g x" by blast hence "(\D. (f has_vector_derivative D) (at x within A)) = (\D. (g has_vector_derivative D) (at x within A))" using assms by (intro ext has_vector_derivative_cong_ev refl assms) simp_all thus ?thesis by (simp add: vector_derivative_def assms) qed lemma differentiable_of_real [simp]: "of_real differentiable at x within A" proof - have "(of_real has_vector_derivative 1) (at x within A)" by (auto intro!: derivative_eq_intros) thus ?thesis by (rule differentiableI_vector) qed lemma higher_deriv_cong_ev: assumes "eventually (\x. f x = g x) (nhds x)" "x = y" shows "(deriv ^^ n) f x = (deriv ^^ n) g y" proof - from assms(1) have "eventually (\x. (deriv ^^ n) f x = (deriv ^^ n) g x) (nhds x)" proof (induction n arbitrary: f g) case (Suc n) from Suc.prems have "eventually (\y. eventually (\z. f z = g z) (nhds y)) (nhds x)" by (simp add: eventually_eventually) hence "eventually (\x. deriv f x = deriv g x) (nhds x)" by eventually_elim (rule deriv_cong_ev, simp_all) thus ?case by (auto intro!: deriv_cong_ev Suc simp: funpow_Suc_right simp del: funpow.simps) qed auto from eventually_nhds_x_imp_x[OF this] assms(2) show ?thesis by simp qed lemma deriv_of_real [simp]: "at x within A \ bot \ vector_derivative of_real (at x within A) = 1" by (auto intro!: vector_derivative_within derivative_eq_intros) lemma deriv_Re [simp]: "deriv Re = (\_. 1)" - by (auto intro!: DERIV_imp_deriv simp: fun_eq_iff) + by simp lemma vector_derivative_of_real_left: assumes "f differentiable at x" shows "vector_derivative (\x. of_real (f x)) (at x) = of_real (deriv f x)" proof - have "vector_derivative (of_real \ f) (at x) = (of_real (deriv f x))" by (subst vector_derivative_chain_at) (simp_all add: scaleR_conv_of_real field_derivative_eq_vector_derivative assms) thus ?thesis by (simp add: o_def) qed lemma vector_derivative_of_real_right: assumes "f field_differentiable at (of_real x)" shows "vector_derivative (\x. f (of_real x)) (at x) = deriv f (of_real x)" proof - have "vector_derivative (f \ of_real) (at x) = deriv f (of_real x)" using assms by (subst vector_derivative_chain_at_general) simp_all thus ?thesis by (simp add: o_def) qed lemma Ln_holomorphic [holomorphic_intros]: assumes "A \ \\<^sub>\\<^sub>0 = {}" shows "Ln holomorphic_on (A :: complex set)" proof (intro holomorphic_onI) fix z assume "z \ A" with assms have "(Ln has_field_derivative inverse z) (at z within A)" by (auto intro!: derivative_eq_intros) thus "Ln field_differentiable at z within A" by (auto simp: field_differentiable_def) qed lemma higher_deriv_Polygamma: assumes "z \ \\<^sub>\\<^sub>0" shows "(deriv ^^ n) (Polygamma m) z = Polygamma (m + n) (z :: 'a :: {real_normed_field,euclidean_space})" proof - have "eventually (\u. (deriv ^^ n) (Polygamma m) u = Polygamma (m + n) u) (nhds z)" proof (induction n) case (Suc n) from Suc.IH have "eventually (\z. eventually (\u. (deriv ^^ n) (Polygamma m) u = Polygamma (m + n) u) (nhds z)) (nhds z)" by (simp add: eventually_eventually) hence "eventually (\z. deriv ((deriv ^^ n) (Polygamma m)) z = deriv (Polygamma (m + n)) z) (nhds z)" by eventually_elim (intro deriv_cong_ev refl) moreover have "eventually (\z. z \ UNIV - \\<^sub>\\<^sub>0) (nhds z)" using assms by (intro eventually_nhds_in_open open_Diff open_UNIV) auto ultimately show ?case by eventually_elim (simp_all add: deriv_Polygamma) qed simp_all thus ?thesis by (rule eventually_nhds_x_imp_x) qed lemma higher_deriv_cmult: assumes "f holomorphic_on A" "x \ A" "open A" shows "(deriv ^^ j) (\x. c * f x) x = c * (deriv ^^ j) f x" using assms proof (induction j arbitrary: f x) case (Suc j f x) have "deriv ((deriv ^^ j) (\x. c * f x)) x = deriv (\x. c * (deriv ^^ j) f x) x" using eventually_nhds_in_open[of A x] assms(2,3) Suc.prems by (intro deriv_cong_ev refl) (auto elim!: eventually_mono simp: Suc.IH) also have "\ = c * deriv ((deriv ^^ j) f) x" using Suc.prems assms(2,3) by (intro deriv_cmult holomorphic_on_imp_differentiable_at holomorphic_higher_deriv) auto finally show ?case by simp qed simp_all lemma higher_deriv_ln_Gamma_complex: assumes "(x::complex) \ \\<^sub>\\<^sub>0" shows "(deriv ^^ j) ln_Gamma x = (if j = 0 then ln_Gamma x else Polygamma (j - 1) x)" proof (cases j) case (Suc j') have "(deriv ^^ j') (deriv ln_Gamma) x = (deriv ^^ j') Digamma x" using eventually_nhds_in_open[of "UNIV - \\<^sub>\\<^sub>0" x] assms by (intro higher_deriv_cong_ev refl) (auto elim!: eventually_mono simp: open_Diff deriv_ln_Gamma_complex) also have "\ = Polygamma j' x" using assms by (subst higher_deriv_Polygamma) (auto elim!: nonpos_Ints_cases simp: complex_nonpos_Reals_iff) finally show ?thesis using Suc by (simp del: funpow.simps add: funpow_Suc_right) qed simp_all lemma higher_deriv_ln_Gamma_real: assumes "(x::real) > 0" shows "(deriv ^^ j) ln_Gamma x = (if j = 0 then ln_Gamma x else Polygamma (j - 1) x)" proof (cases j) case (Suc j') have "(deriv ^^ j') (deriv ln_Gamma) x = (deriv ^^ j') Digamma x" using eventually_nhds_in_open[of "{0<..}" x] assms by (intro higher_deriv_cong_ev refl) (auto elim!: eventually_mono simp: open_Diff deriv_ln_Gamma_real) also have "\ = Polygamma j' x" using assms by (subst higher_deriv_Polygamma) (auto elim!: nonpos_Ints_cases simp: complex_nonpos_Reals_iff) finally show ?thesis using Suc by (simp del: funpow.simps add: funpow_Suc_right) qed simp_all lemma higher_deriv_ln_Gamma_complex_of_real: assumes "(x :: real) > 0" shows "(deriv ^^ j) ln_Gamma (complex_of_real x) = of_real ((deriv ^^ j) ln_Gamma x)" using assms by (auto simp: higher_deriv_ln_Gamma_real higher_deriv_ln_Gamma_complex ln_Gamma_complex_of_real Polygamma_of_real) (* END TODO *) (* TODO: could be automated with Laurent series expansions in the future *) lemma stirling_limit_aux1: "((\y. Ln (1 + z * of_real y) / of_real y) \ z) (at_right 0)" for z :: complex proof (cases "z = 0") case True then show ?thesis by simp next case False have "((\y. ln (1 + z * of_real y)) has_vector_derivative 1 * z) (at 0)" by (rule has_vector_derivative_real_field) (auto intro!: derivative_eq_intros) then have "(\y. (Ln (1 + z * of_real y) - of_real y * z) / of_real \y\) \0\ 0" by (auto simp add: has_vector_derivative_def has_derivative_def netlimit_at scaleR_conv_of_real field_simps) then have "((\y. (Ln (1 + z * of_real y) - of_real y * z) / of_real \y\) \ 0) (at_right 0)" by (rule filterlim_mono[OF _ _ at_le]) simp_all also have "?this \ ((\y. Ln (1 + z * of_real y) / (of_real y) - z) \ 0) (at_right 0)" using eventually_at_right_less[of "0::real"] by (intro filterlim_cong refl) (auto elim!: eventually_mono simp: field_simps) finally show ?thesis by (simp only: LIM_zero_iff) qed lemma stirling_limit_aux2: "((\y. y * Ln (1 + z / of_real y)) \ z) at_top" for z :: complex using stirling_limit_aux1[of z] by (subst filterlim_at_top_to_right) (simp add: field_simps) lemma Union_atLeastAtMost: assumes "N > 0" shows "(\n\{0.. {0..real N}" thus "x \ (\n\{0.. 0" "x < real N" by simp_all hence "x \ real (nat \x\)" "x \ real (nat \x\ + 1)" by linarith+ moreover from x have "nat \x\ < N" by linarith ultimately have "\n\{0.. {real n..real (n + 1)}" by (intro bexI[of _ "nat \x\"]) simp_all thus ?thesis by blast qed qed auto subsection \Cones in the complex plane\ definition complex_cone :: "real \ real \ complex set" where "complex_cone a b = {z. \y\{a..b}. z = rcis (norm z) y}" abbreviation complex_cone' :: "real \ complex set" where "complex_cone' a \ complex_cone (-a) a" lemma zero_in_complex_cone [simp, intro]: "a \ b \ 0 \ complex_cone a b" by (auto simp: complex_cone_def) lemma complex_coneE: assumes "z \ complex_cone a b" obtains r \ where "r \ 0" "\ \ {a..b}" "z = rcis r \" proof - from assms obtain y where "y \ {a..b}" "z = rcis (norm z) y" unfolding complex_cone_def by auto thus ?thesis using that[of "norm z" y] by auto qed lemma arg_cis [simp]: assumes "x \ {-pi<..pi}" - shows "arg (cis x) = x" + shows "Arg (cis x) = x" using assms by (intro arg_unique) auto lemma arg_mult_of_real_left [simp]: assumes "r > 0" - shows "arg (of_real r * z) = arg z" + shows "Arg (of_real r * z) = Arg z" proof (cases "z = 0") case False thus ?thesis - using arg_bounded[of z] assms - by (intro arg_unique) (auto simp: sgn_mult sgn_of_real cis_arg) + using Arg_bounded[of z] assms + by (intro arg_unique) (auto simp: sgn_mult sgn_of_real cis_Arg) qed auto lemma arg_mult_of_real_right [simp]: assumes "r > 0" - shows "arg (z * of_real r) = arg z" + shows "Arg (z * of_real r) = Arg z" by (subst mult.commute, subst arg_mult_of_real_left) (simp_all add: assms) lemma arg_rcis [simp]: assumes "x \ {-pi<..pi}" "r > 0" - shows "arg (rcis r x) = x" + shows "Arg (rcis r x) = x" using assms by (simp add: rcis_def) lemma rcis_in_complex_cone [intro]: assumes "\ \ {a..b}" "r \ 0" shows "rcis r \ \ complex_cone a b" using assms by (auto simp: complex_cone_def) lemma arg_imp_in_complex_cone: - assumes "arg z \ {a..b}" + assumes "Arg z \ {a..b}" shows "z \ complex_cone a b" proof - - have "z = rcis (norm z) (arg z)" - by (simp add: rcis_cmod_arg) + have "z = rcis (norm z) (Arg z)" + by (simp add: rcis_cmod_Arg) also have "\ \ complex_cone a b" using assms by auto finally show ?thesis . qed lemma complex_cone_altdef: assumes "-pi < a" "a \ b" "b \ pi" - shows "complex_cone a b = insert 0 {z. arg z \ {a..b}}" + shows "complex_cone a b = insert 0 {z. Arg z \ {a..b}}" proof (intro equalityI subsetI) fix z assume "z \ complex_cone a b" then obtain r \ where *: "r \ 0" "\ \ {a..b}" "z = rcis r \" by (auto elim: complex_coneE) - have "arg z \ {a..b}" if [simp]: "z \ 0" + have "Arg z \ {a..b}" if [simp]: "z \ 0" proof - have "r > 0" using that * by (subst (asm) *) auto hence "\ \ {a..b}" using *(1,2) assms by (auto simp: *(1)) moreover from assms *(2) have "\ \ {-pi<..pi}" by auto ultimately show ?thesis using *(3) \r > 0\ by (subst *) auto qed - thus "z \ insert 0 {z. arg z \ {a..b}}" + thus "z \ insert 0 {z. Arg z \ {a..b}}" by auto qed (use assms in \auto intro: arg_imp_in_complex_cone\) lemma nonneg_of_real_in_complex_cone [simp, intro]: assumes "x \ 0" "a \ 0" "0 \ b" shows "of_real x \ complex_cone a b" proof - from assms have "rcis x 0 \ complex_cone a b" by (intro rcis_in_complex_cone) auto thus ?thesis by simp qed lemma one_in_complex_cone [simp, intro]: "a \ 0 \ 0 \ b \ 1 \ complex_cone a b" using nonneg_of_real_in_complex_cone[of 1] by (simp del: nonneg_of_real_in_complex_cone) lemma of_nat_in_complex_cone [simp, intro]: "a \ 0 \ 0 \ b \ of_nat n \ complex_cone a b" using nonneg_of_real_in_complex_cone[of "real n"] by (simp del: nonneg_of_real_in_complex_cone) subsection \Another integral representation of the Beta function\ lemma complex_cone_inter_nonpos_Reals: assumes "-pi < a" "a \ b" "b < pi" shows "complex_cone a b \ \\<^sub>\\<^sub>0 = {0}" proof (safe elim!: nonpos_Reals_cases) fix x :: real assume "complex_of_real x \ complex_cone a b" "x \ 0" hence "\(x < 0)" using assms by (intro notI) (auto simp: complex_cone_altdef) with \x \ 0\ show "complex_of_real x = 0" by auto qed (use assms in auto) theorem assumes a: "a > 0" and b: "b > (0 :: real)" shows has_integral_Beta_real': "((\u. u powr (b - 1) / (1 + u) powr (a + b)) has_integral Beta a b) {0<..}" and Beta_conv_nn_integral: "Beta a b = (\\<^sup>+u. ennreal (indicator {0<..} u * u powr (b - 1) / (1 + u) powr (a + b)) \lborel)" proof - define I where "I = (\\<^sup>+u. ennreal (indicator {0<..} u * u powr (b - 1) / (1 + u) powr (a + b)) \lborel)" have "Gamma (a + b) > 0" "Beta a b > 0" using assms by (simp_all add: add_pos_pos Beta_def) from a b have "ennreal (Gamma a * Gamma b) = (\\<^sup>+ t. ennreal (indicator {0..} t * t powr (a - 1) / exp t) \lborel) * (\\<^sup>+ t. ennreal (indicator {0..} t * t powr (b - 1) / exp t) \lborel)" by (subst ennreal_mult') (simp_all add: Gamma_conv_nn_integral_real) also have "\ = (\\<^sup>+t. \\<^sup>+u. ennreal (indicator {0..} t * t powr (a - 1) / exp t) * ennreal (indicator {0..} u * u powr (b - 1) / exp u) \lborel \lborel)" by (simp add: nn_integral_cmult nn_integral_multc) also have "\ = (\\<^sup>+t. indicator {0<..} t * (\\<^sup>+u. indicator {0..} u * t powr (a - 1) * u powr (b - 1) / exp (t + u) \lborel) \lborel)" by (intro nn_integral_cong_AE AE_I[of _ _ "{0}"]) (auto simp: indicator_def divide_ennreal ennreal_mult' [symmetric] exp_add mult_ac) also have "\ = (\\<^sup>+t. indicator {0<..} t * (\\<^sup>+u. indicator {0..} u * t powr (a - 1) * u powr (b - 1) / exp (t + u) \(density (distr lborel borel ((*) t)) (\x. ennreal \t\))) \lborel)" by (intro nn_integral_cong mult_indicator_cong, subst lborel_distr_mult' [symmetric]) auto also have "\ = (\\<^sup>+(t::real). indicator {0<..} t * (\\<^sup>+u. indicator {0..} (u * t) * t powr a * (u * t) powr (b - 1) / exp (t + t * u) \lborel) \lborel)" by (intro nn_integral_cong mult_indicator_cong) (auto simp: nn_integral_density nn_integral_distr algebra_simps powr_diff simp flip: ennreal_mult) also have "\ = (\\<^sup>+(t::real). \\<^sup>+u. indicator ({0<..}\{0..}) (t, u) * t powr a * (u * t) powr (b - 1) / exp (t * (u + 1)) \lborel \lborel)" by (subst nn_integral_cmult [symmetric], simp, intro nn_integral_cong) (auto simp: indicator_def zero_le_mult_iff algebra_simps) also have "\ = (\\<^sup>+(t::real). \\<^sup>+u. indicator ({0<..}\{0..}) (t, u) * t powr (a + b - 1) * u powr (b - 1) / exp (t * (u + 1)) \lborel \lborel)" by (intro nn_integral_cong) (auto simp: powr_add powr_diff indicator_def powr_mult field_simps) also have "\ = (\\<^sup>+(u::real). \\<^sup>+t. indicator ({0<..}\{0..}) (t, u) * t powr (a + b - 1) * u powr (b - 1) / exp (t * (u + 1)) \lborel \lborel)" by (rule lborel_pair.Fubini') auto also have "\ = (\\<^sup>+(u::real). indicator {0..} u * (\\<^sup>+t. indicator {0<..} t * t powr (a + b - 1) * u powr (b - 1) / exp (t * (u + 1)) \lborel) \lborel)" by (intro nn_integral_cong mult_indicator_cong) (auto simp: indicator_def) also have "\ = (\\<^sup>+(u::real). indicator {0<..} u * (\\<^sup>+t. indicator {0<..} t * t powr (a + b - 1) * u powr (b - 1) / exp (t * (u + 1)) \lborel) \lborel)" by (intro nn_integral_cong_AE AE_I[of _ _ "{0}"]) (auto simp: indicator_def) also have "\ = (\\<^sup>+(u::real). indicator {0<..} u * (\\<^sup>+t. indicator {0<..} t * t powr (a + b - 1) * u powr (b - 1) / exp (t * (u + 1)) \(density (distr lborel borel ((*) (1/(1+u)))) (\x. ennreal \1/(1+u)\))) \lborel)" by (intro nn_integral_cong mult_indicator_cong, subst lborel_distr_mult' [symmetric]) auto also have "\ = (\\<^sup>+(u::real). indicator {0<..} u * (\\<^sup>+t. ennreal (1 / (u + 1)) * ennreal (indicator {0<..} (t / (u + 1)) * (t / (1+u)) powr (a + b - 1) * u powr (b - 1) / exp t) \lborel) \lborel)" by (intro nn_integral_cong mult_indicator_cong) (auto simp: nn_integral_distr nn_integral_density add_ac) also have "\ = (\\<^sup>+u. \\<^sup>+t. indicator ({0<..}\{0<..}) (u, t) * 1/(u+1) * (t / (u+1)) powr (a + b - 1) * u powr (b - 1) / exp t \lborel \lborel)" by (subst nn_integral_cmult [symmetric], simp, intro nn_integral_cong) (auto simp: indicator_def field_simps divide_ennreal simp flip: ennreal_mult ennreal_mult') also have "\ = (\\<^sup>+u. \\<^sup>+t. ennreal (indicator {0<..} u * u powr (b - 1) / (1 + u) powr (a + b)) * ennreal (indicator {0<..} t * t powr (a + b - 1) / exp t) \lborel \lborel)" by (intro nn_integral_cong) (auto simp: indicator_def powr_add powr_diff powr_divide powr_minus divide_simps add_ac simp flip: ennreal_mult) also have "\ = I * (\\<^sup>+t. indicator {0<..} t * t powr (a + b - 1) / exp t \lborel)" by (simp add: nn_integral_cmult nn_integral_multc I_def) also have "(\\<^sup>+t. indicator {0<..} t * t powr (a + b - 1) / exp t \lborel) = ennreal (Gamma (a + b))" using assms by (subst Gamma_conv_nn_integral_real) (auto intro!: nn_integral_cong_AE[OF AE_I[of _ _ "{0}"]] simp: indicator_def split: if_splits) finally have "ennreal (Gamma a * Gamma b) = I * ennreal (Gamma (a + b))" . hence "ennreal (Gamma a * Gamma b) / ennreal (Gamma (a + b)) = I * ennreal (Gamma (a + b)) / ennreal (Gamma (a + b))" by simp also have "\ = I" using \Gamma (a + b) > 0\ by (intro ennreal_mult_divide_eq) (auto simp: ) also have "ennreal (Gamma a * Gamma b) / ennreal (Gamma (a + b)) = ennreal (Gamma a * Gamma b / Gamma (a + b))" using assms by (intro divide_ennreal) auto also have "\ = ennreal (Beta a b)" by (simp add: Beta_def) finally show *: "ennreal (Beta a b) = I" . define f where "f = (\u. u powr (b - 1) / (1 + u) powr (a + b))" have "((\u. indicator {0<..} u * f u) has_integral Beta a b) UNIV" using * \Beta a b > 0\ by (subst has_integral_iff_nn_integral_lebesgue) (auto simp: f_def measurable_completion nn_integral_completion I_def mult_ac) also have "(\u. indicator {0<..} u * f u) = (\u. if u \ {0<..} then f u else 0)" by (auto simp: fun_eq_iff) also have "(\ has_integral Beta a b) UNIV \ (f has_integral Beta a b) {0<..}" by (rule has_integral_restrict_UNIV) finally show \ by (simp add: f_def) qed lemma has_integral_Beta2: fixes a :: real assumes "a < -1/2" shows "((\x. (1 + x ^ 2) powr a) has_integral Beta (- a - 1 / 2) (1 / 2) / 2) {0<..}" proof - define f where "f = (\u. u powr (-1/2) / (1 + u) powr (-a))" define C where "C = Beta (-a-1/2) (1/2)" have I: "(f has_integral C) {0<..}" using has_integral_Beta_real'[of "-a-1/2" "1/2"] assms by (simp_all add: diff_divide_distrib f_def C_def) define g where "g = (\x. x ^ 2 :: real)" have bij: "bij_betw g {0<..} {0<..}" by (intro bij_betwI[of _ _ _ sqrt]) (auto simp: g_def) have "(f absolutely_integrable_on g ` {0<..} \ integral (g ` {0<..}) f = C)" using I bij by (simp add: bij_betw_def has_integral_iff absolutely_integrable_on_def f_def) also have "?this \ ((\x. \2 * x\ *\<^sub>R f (g x)) absolutely_integrable_on {0<..} \ integral {0<..} (\x. \2 * x\ *\<^sub>R f (g x)) = C)" using bij by (intro has_absolute_integral_change_of_variables_1' [symmetric]) (auto intro!: derivative_eq_intros simp: g_def bij_betw_def) finally have "((\x. \2 * x\ * f (g x)) has_integral C) {0<..}" by (simp add: absolutely_integrable_on_def f_def has_integral_iff) also have "?this \ ((\x::real. 2 * (1 + x\<^sup>2) powr a) has_integral C) {0<..}" by (intro has_integral_cong) (auto simp: f_def g_def powr_def exp_minus ln_realpow field_simps) finally have "((\x::real. 1/2 * (2 * (1 + x\<^sup>2) powr a)) has_integral 1/2 * C) {0<..}" by (intro has_integral_mult_right) thus ?thesis by (simp add: C_def) qed lemma has_integral_Beta3: fixes a b :: real assumes "a < -1/2" and "b > 0" shows "((\x. (b + x ^ 2) powr a) has_integral Beta (-a - 1/2) (1/2) / 2 * b powr (a + 1/2)) {0<..}" proof - define C where "C = Beta (- a - 1 / 2) (1 / 2) / 2" have int: "nn_integral lborel (\x. indicator {0<..} x * (1 + x ^ 2) powr a) = C" using nn_integral_has_integral_lebesgue[OF _ has_integral_Beta2[OF assms(1)]] by (auto simp: C_def) have "nn_integral lborel (\x. indicator {0<..} x * (b + x ^ 2) powr a) = (\\<^sup>+x. ennreal (indicat_real {0<..} (x * sqrt b) * (b + (x * sqrt b)\<^sup>2) powr a * sqrt b) \lborel)" using assms by (subst lborel_distr_mult'[of "sqrt b"]) (auto simp: nn_integral_density nn_integral_distr mult_ac simp flip: ennreal_mult) also have "\ = (\\<^sup>+x. ennreal (indicat_real {0<..} x * (b * (1 + x ^ 2)) powr a * sqrt b) \lborel)" using assms by (intro nn_integral_cong) (auto simp: indicator_def field_simps zero_less_mult_iff) also have "\ = (\\<^sup>+x. ennreal (indicat_real {0<..} x * b powr (a + 1/2) * (1 + x ^ 2) powr a) \lborel)" using assms by (intro nn_integral_cong) (auto simp: indicator_def powr_add powr_half_sqrt powr_mult) also have "\ = b powr (a + 1/2) * (\\<^sup>+x. ennreal (indicat_real {0<..} x * (1 + x ^ 2) powr a) \lborel)" using assms by (subst nn_integral_cmult [symmetric]) (simp_all add: mult_ac flip: ennreal_mult) also have "(\\<^sup>+x. ennreal (indicat_real {0<..} x * (1 + x ^ 2) powr a) \lborel) = C" using int by simp also have "ennreal (b powr (a + 1/2)) * ennreal C = ennreal (C * b powr (a + 1/2))" using assms by (subst ennreal_mult) (auto simp: C_def mult_ac Beta_def) finally have *: "(\\<^sup>+ x. ennreal (indicat_real {0<..} x * (b + x\<^sup>2) powr a) \lborel) = \" . hence "((\x. indicator {0<..} x * (b + x^2) powr a) has_integral C * b powr (a + 1/2)) UNIV" using assms by (subst has_integral_iff_nn_integral_lebesgue) (auto simp: C_def measurable_completion nn_integral_completion Beta_def) also have "(\x. indicator {0<..} x * (b + x^2) powr a) = (\x. if x \ {0<..} then (b + x^2) powr a else 0)" by (auto simp: fun_eq_iff) finally show ?thesis by (subst (asm) has_integral_restrict_UNIV) (auto simp: C_def) qed subsection \Asymptotics of the real $\log\Gamma$ function and its derivatives\ text \ This is the error term that occurs in the expansion of @{term ln_Gamma}. It can be shown to be of order $O(s^{-n})$. \ definition stirling_integral :: "nat \ 'a :: {real_normed_div_algebra, banach} \ 'a" where "stirling_integral n s = lim (\N. integral {0..N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n))" context fixes s :: complex assumes s: "s \ \\<^sub>\\<^sub>0" fixes approx :: "nat \ complex" defines "approx \ (\N. (\n = 1.. \\\ ln_Gamma s\\ (ln_Gamma (of_nat N) - ln (2 * pi / of_nat N) / 2 - of_nat N * ln (of_nat N) + of_nat N) - \ \\\ 0\\ s * (harm (N - 1) - ln (of_nat (N - 1)) - euler_mascheroni) + \ \\\ 0\\ s * (ln (of_nat N + s) - ln (of_nat (N - 1))) - \ \\\ 0\\ (1/2) * (ln (of_nat N + s) - ln (of_nat N)) + \ \\\ 0\\ of_nat N * (ln (of_nat N + s) - ln (of_nat N)) - \ \\\ s\\ (s - 1/2) * ln s - ln (2 * pi) / 2)" begin qualified lemma assumes N: "N > 0" shows integrable_pbernpoly_1: "(\x. of_real (-pbernpoly 1 x) / (of_real x + s)) integrable_on {0..real N}" and integral_pbernpoly_1_aux: "integral {0..real N} (\x. -of_real (pbernpoly 1 x) / (of_real x + s)) = approx N" and has_integral_pbernpoly_1: "((\x. pbernpoly 1 x /(x + s)) has_integral (\mx. -pbernpoly 1 x / (x + s)) has_integral (of_nat n + 1/2 + s) * (ln (of_nat (n + 1) + s) - ln (of_nat n + s)) - 1) {of_nat n..of_nat (n + 1)}" for n proof (rule has_integral_spike) have "((\x. (of_nat n + 1/2 + s) * (1 / (x + s)) - 1) has_integral (of_nat n + 1/2 + s) * (ln (of_real (real (n + 1)) + s) - ln (of_real (real n) + s)) - 1) {of_nat n..of_nat (n + 1)}" using s has_integral_const_real[of 1 "of_nat n" "of_nat (n + 1)"] by (intro has_integral_diff has_integral_mult_right fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros has_vector_derivative_real_field simp: has_field_derivative_iff_has_vector_derivative [symmetric] field_simps complex_nonpos_Reals_iff) thus "((\x. (of_nat n + 1/2 + s) * (1 / (x + s)) - 1) has_integral (of_nat n + 1/2 + s) * (ln (of_nat (n + 1) + s) - ln (of_nat n + s)) - 1) {of_nat n..of_nat (n + 1)}" by simp show "-pbernpoly 1 x / (x + s) = (of_nat n + 1/2 + s) * (1 / (x + s)) - 1" if "x \ {of_nat n..of_nat (n + 1)} - {of_nat (n + 1)}" for x proof - have x: "x \ real n" "x < real (n + 1)" using that by simp_all hence "floor x = int n" by linarith moreover from s x have "complex_of_real x \ -s" by (auto simp add: complex_eq_iff complex_nonpos_Reals_iff simp del: of_nat_Suc) ultimately show "-pbernpoly 1 x / (x + s) = (of_nat n + 1/2 + s) * (1 / (x + s)) - 1" by (auto simp: pbernpoly_def bernpoly_def frac_def divide_simps add_eq_0_iff2) qed qed simp_all hence *: "\I. I\?A \ ((\x. -pbernpoly 1 x / (x + s)) has_integral (Inf I + 1/2 + s) * (ln (Inf I + 1 + s) - ln (Inf I + s)) - 1) I" by (auto simp: add_ac) have "((\x. - pbernpoly 1 x / (x + s)) has_integral (\I\?A. (Inf I + 1 / 2 + s) * (ln (Inf I + 1 + s) - ln (Inf I + s)) - 1)) (\n\{0..x. - pbernpoly 1 x / (x + s)) has_integral ?i) {0..real N}" by (subst has_integral_spike_set_eq) (use Union_atLeastAtMost assms in \auto simp: intro!: empty_imp_negligible\) hence "(\x. - pbernpoly 1 x / (x + s)) integrable_on {0..real N}" and integral: "integral {0..real N} (\x. - pbernpoly 1 x / (x + s)) = ?i" by (simp_all add: has_integral_iff) show "(\x. - pbernpoly 1 x / (x + s)) integrable_on {0..real N}" by fact note has_integral_neg[OF has_integral] also have "-?i = (\xx. of_real (pbernpoly 1 x) / (of_real x + s)) has_integral \) {0..real N}" by simp note integral also have "?i = (\nnnn=1.. = (\n=1..n=1..n=1..n=1.. - (\n = 1..n=1.. 0" for x by (auto simp: complex_nonpos_Reals_iff complex_eq_iff) hence "(\n=1..n=1.. = ln (fact (N - 1)) + (\n=1..n=1..n=1..n=1..n = 1..n = 1..x. -of_real (pbernpoly 1 x) / (of_real x + s)) = \" by simp qed lemma integrable_ln_Gamma_aux: shows "(\x. of_real (pbernpoly n x) / (of_real x + s) ^ n) integrable_on {0..real N}" proof (cases "n = 1") case True with s show ?thesis using integrable_neg[OF integrable_pbernpoly_1[of N]] by (cases "N = 0") (simp_all add: integrable_negligible) next case False from s have "of_real x + s \ 0" if "x \ 0" for x using that by (auto simp: complex_eq_iff add_eq_0_iff2 complex_nonpos_Reals_iff) with False s show ?thesis by (auto intro!: integrable_continuous_real continuous_intros) qed text \ This following proof is based on ``Rudiments of the theory of the gamma function'' by Bruce Berndt~\cite{berndt}. \ qualified lemma integral_pbernpoly_1: "(\N. integral {0..real N} (\x. pbernpoly 1 x / (x + s))) \ -ln_Gamma s - s + (s - 1 / 2) * ln s + ln (2 * pi) / 2" proof - have neq: "s + of_real x \ 0" if "x \ 0" for x :: real using that s by (auto simp: complex_eq_iff complex_nonpos_Reals_iff) have "(approx \ ln_Gamma s - 0 - 0 + 0 - 0 + s - (s - 1/2) * ln s - ln (2 * pi) / 2) at_top" unfolding approx_def proof (intro tendsto_add tendsto_diff) from s have s': "s \ \\<^sub>\\<^sub>0" by (auto simp: complex_nonpos_Reals_iff elim!: nonpos_Ints_cases) have "(\n. \i=1.. ln_Gamma s + euler_mascheroni * s + ln s" (is "?f \ _") using ln_Gamma_series'_aux[OF s'] unfolding sums_def by (subst filterlim_sequentially_Suc [symmetric], subst (asm) sum.atLeast1_atMost_eq [symmetric]) (simp add: atLeastLessThanSuc_atLeastAtMost) thus "((\n. ?f n - (euler_mascheroni * s + ln s)) \ ln_Gamma s) at_top" by (auto intro: tendsto_eq_intros) next show "(\x. complex_of_real (ln_Gamma (real x) - ln (2 * pi / real x) / 2 - real x * ln (real x) + real x)) \ 0" proof (intro tendsto_of_real_0_I filterlim_compose[OF tendsto_sandwich filterlim_real_sequentially]) show "eventually (\x::real. ln_Gamma x - ln (2 * pi / x) / 2 - x * ln x + x \ 0) at_top" using eventually_ge_at_top[of "1::real"] by eventually_elim (insert ln_Gamma_bounds(1), simp add: algebra_simps) show "eventually (\x::real. ln_Gamma x - ln (2 * pi / x) / 2 - x * ln x + x \ 1 / 12 * inverse x) at_top" using eventually_ge_at_top[of "1::real"] by eventually_elim (insert ln_Gamma_bounds(2), simp add: field_simps) show "((\x::real. 1 / 12 * inverse x) \ 0) at_top" by (intro tendsto_mult_right_zero tendsto_inverse_0_at_top filterlim_ident) qed simp_all next have "(\x. s * of_real (harm (x - 1) - ln (real (x - 1)) - euler_mascheroni)) \ s * of_real (euler_mascheroni - euler_mascheroni)" by (subst filterlim_sequentially_Suc [symmetric], intro tendsto_intros) (insert euler_mascheroni_LIMSEQ, simp_all) also have "?this \ (\x. s * (harm (x - 1) - ln (of_nat (x - 1)) - euler_mascheroni)) \ 0" by (intro filterlim_cong refl eventually_mono[OF eventually_gt_at_top[of "1::nat"]]) (auto simp: Ln_of_nat of_real_harm) finally show "(\x. s * (harm (x - 1) - ln (of_nat (x - 1)) - euler_mascheroni)) \ 0" . next have "((\x. ln (1 + (s + 1) / of_real x)) \ ln (1 + 0)) at_top" (is ?P) by (intro tendsto_intros tendsto_divide_0[OF tendsto_const]) (simp_all add: filterlim_ident filterlim_at_infinity_conv_norm_at_top filterlim_abs_real) also have "ln (of_real (x + 1) + s) - ln (complex_of_real x) = ln (1 + (s + 1) / of_real x)" if "x > 1" for x using that s using Ln_divide_of_real[of x "of_real (x + 1) + s", symmetric] neq[of "x+1"] by (simp add: field_simps Ln_of_real) hence "?P \ ((\x. ln (of_real (x + 1) + s) - ln (of_real x)) \ 0) at_top" by (intro filterlim_cong refl) (auto intro: eventually_mono[OF eventually_gt_at_top[of "1::real"]]) finally have "((\n. ln (of_real (real n + 1) + s) - ln (of_real (real n))) \ 0) at_top" by (rule filterlim_compose[OF _ filterlim_real_sequentially]) hence "((\n. ln (of_nat n + s) - ln (of_nat (n - 1))) \ 0) at_top" by (subst filterlim_sequentially_Suc [symmetric]) (simp add: add_ac) thus "(\x. s * (ln (of_nat x + s) - ln (of_nat (x - 1)))) \ 0" by (rule tendsto_mult_right_zero) next have "((\x. ln (1 + s / of_real x)) \ ln (1 + 0)) at_top" (is ?P) by (intro tendsto_intros tendsto_divide_0[OF tendsto_const]) (simp_all add: filterlim_ident filterlim_at_infinity_conv_norm_at_top filterlim_abs_real) also have "ln (of_real x + s) - ln (of_real x) = ln (1 + s / of_real x)" if "x > 0" for x using Ln_divide_of_real[of x "of_real x + s"] neq[of x] that by (auto simp: field_simps Ln_of_real) hence "?P \ ((\x. ln (of_real x + s) - ln (of_real x)) \ 0) at_top" using s by (intro filterlim_cong refl) (auto intro: eventually_mono [OF eventually_gt_at_top[of "1::real"]]) finally have "(\x. (1/2) * (ln (of_real (real x) + s) - ln (of_real (real x)))) \ 0" by (rule tendsto_mult_right_zero[OF filterlim_compose[OF _ filterlim_real_sequentially]]) thus "(\x. (1/2) * (ln (of_nat x + s) - ln (of_nat x))) \ 0" by simp next have "((\x. x * (ln (1 + s / of_real x))) \ s) at_top" (is ?P) by (rule stirling_limit_aux2) also have "ln (1 + s / of_real x) = ln (of_real x + s) - ln (of_real x)" if "x > 1" for x using that s Ln_divide_of_real [of x "of_real x + s", symmetric] neq[of x] by (auto simp: Ln_of_real field_simps) hence "?P \ ((\x. of_real x * (ln (of_real x + s) - ln (of_real x))) \ s) at_top" by (intro filterlim_cong refl) (auto intro: eventually_mono[OF eventually_gt_at_top[of "1::real"]]) finally have "(\n. of_real (real n) * (ln (of_real (real n) + s) - ln (of_real (real n)))) \ s" by (rule filterlim_compose[OF _ filterlim_real_sequentially]) thus "(\n. of_nat n * (ln (of_nat n + s) - ln (of_nat n))) \ s" by simp qed simp_all also have "?this \ ((\N. integral {0..real N} (\x. -pbernpoly 1 x / (x + s))) \ ln_Gamma s + s - (s - 1/2) * ln s - ln (2 * pi) / 2) at_top" using integral_pbernpoly_1_aux by (intro filterlim_cong refl) (auto intro: eventually_mono[OF eventually_gt_at_top[of "0::nat"]]) also have "(\N. integral {0..real N} (\x. -pbernpoly 1 x / (x + s))) = (\N. -integral {0..real N} (\x. pbernpoly 1 x / (x + s)))" by (simp add: fun_eq_iff) finally show ?thesis by (simp add: tendsto_minus_cancel_left [symmetric] algebra_simps) qed qualified lemma pbernpoly_integral_conv_pbernpoly_integral_Suc: assumes "n \ 1" shows "integral {0..real N} (\x. pbernpoly n x / (x + s) ^ n) = of_real (pbernpoly (Suc n) (real N)) / (of_nat (Suc n) * (s + of_nat N) ^ n) - of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n) + of_nat n / of_nat (Suc n) * integral {0..real N} (\x. of_real (pbernpoly (Suc n) x) / (of_real x + s) ^ Suc n)" proof - note [derivative_intros] = has_field_derivative_pbernpoly_Suc' define I where "I = -of_real (pbernpoly (Suc n) (of_nat N)) / (of_nat (Suc n) * (of_nat N + s) ^ n) + of_real (bernoulli (Suc n) / real (Suc n)) / s ^ n + integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)" have "((\x. (-of_nat n * inverse (of_real x + s) ^ Suc n) * (of_real (pbernpoly (Suc n) x) / (of_nat (Suc n)))) has_integral -I) {0..real N}" proof (rule integration_by_parts_interior_strong[OF bounded_bilinear_mult]) fix x :: real assume x: "x \ {0<.. \" proof assume "x \ \" then obtain n where "x = of_int n" by (auto elim!: Ints_cases) with x have x': "x = of_nat (nat n)" by simp from x show False by (auto simp: x') qed hence "((\x. of_real (pbernpoly (Suc n) x / of_nat (Suc n))) has_vector_derivative complex_of_real (pbernpoly n x)) (at x)" by (intro has_vector_derivative_of_real) (auto intro!: derivative_eq_intros) thus "((\x. of_real (pbernpoly (Suc n) x) / of_nat (Suc n)) has_vector_derivative complex_of_real (pbernpoly n x)) (at x)" by simp from x s have "complex_of_real x + s \ 0" by (auto simp: complex_eq_iff complex_nonpos_Reals_iff) thus "((\x. inverse (of_real x + s) ^ n) has_vector_derivative - of_nat n * inverse (of_real x + s) ^ Suc n) (at x)" using x s assms by (auto intro!: derivative_eq_intros has_vector_derivative_real_field simp: divide_simps power_add [symmetric] simp del: power_Suc) next have "complex_of_real x + s \ 0" if "x \ 0" for x using that s by (auto simp: complex_eq_iff complex_nonpos_Reals_iff) thus "continuous_on {0..real N} (\x. inverse (of_real x + s) ^ n)" "continuous_on {0..real N} (\x. complex_of_real (pbernpoly (Suc n) x) / of_nat (Suc n))" using assms s by (auto intro!: continuous_intros simp del: of_nat_Suc) next have "((\x. inverse (of_real x + s) ^ n * of_real (pbernpoly n x)) has_integral pbernpoly (Suc n) (of_nat N) / (of_nat (Suc n) * (of_nat N + s) ^ n) - of_real (bernoulli (Suc n) / real (Suc n)) / s ^ n - -I) {0..real N}" using integrable_ln_Gamma_aux[of n N] assms by (auto simp: I_def has_integral_integral divide_simps) thus "((\x. inverse (of_real x + s) ^ n * of_real (pbernpoly n x)) has_integral inverse (of_real (real N) + s) ^ n * (of_real (pbernpoly (Suc n) (real N)) / of_nat (Suc n)) - inverse (of_real 0 + s) ^ n * (of_real (pbernpoly (Suc n) 0) / of_nat (Suc n)) - - I) {0..real N}" by (simp_all add: field_simps) qed simp_all also have "(\x. - of_nat n * inverse (of_real x + s) ^ Suc n * (of_real (pbernpoly (Suc n) x) / of_nat (Suc n))) = (\x. - (of_nat n / of_nat (Suc n) * of_real (pbernpoly (Suc n) x) / (of_real x + s) ^ Suc n))" by (simp add: divide_simps fun_eq_iff) finally have "((\x. - (of_nat n / of_nat (Suc n) * of_real (pbernpoly (Suc n) x) / (of_real x + s) ^ Suc n)) has_integral - I) {0..real N}" . from has_integral_neg[OF this] show ?thesis by (auto simp add: I_def has_integral_iff algebra_simps integral_mult_right [symmetric] simp del: power_Suc of_nat_Suc ) qed lemma pbernpoly_over_power_tendsto_0: assumes "n > 0" shows "(\x. of_real (pbernpoly (Suc n) (real x)) / (of_nat (Suc n) * (s + of_nat x) ^ n)) \ 0" proof - from s have neq: "s + of_nat n \ 0" for n by (auto simp: complex_eq_iff complex_nonpos_Reals_iff) from bounded_pbernpoly[of "Suc n"] guess c . note c = this have "eventually (\x. real x + Re s > 0) at_top" by real_asymp hence "eventually (\x. norm (of_real (pbernpoly (Suc n) (real x)) / (of_nat (Suc n) * (s + of_nat x) ^ n)) \ (c / real (Suc n)) / (real x + Re s) ^ n) at_top" using eventually_gt_at_top[of "0::nat"] proof eventually_elim case (elim x) have "norm (of_real (pbernpoly (Suc n) (real x)) / (of_nat (Suc n) * (s + of_nat x) ^ n)) \ (c / real (Suc n)) / norm (s + of_nat x) ^ n" (is "_ \ ?rhs") using c[of x] by (auto simp: norm_divide norm_mult norm_power neq field_simps simp del: of_nat_Suc) also have "(real x + Re s) \ cmod (s + of_nat x)" using complex_Re_le_cmod[of "s + of_nat x"] s by (auto simp add: complex_nonpos_Reals_iff) hence "?rhs \ (c / real (Suc n)) / (real x + Re s) ^ n" using s elim c[of 0] neq[of x] by (intro divide_left_mono power_mono mult_pos_pos divide_nonneg_pos zero_less_power) auto finally show ?case . qed moreover have "(\x. (c / real (Suc n)) / (real x + Re s) ^ n) \ 0" using \n > 0\ by real_asymp ultimately show ?thesis by (rule Lim_null_comparison) qed lemma convergent_stirling_integral: assumes "n > 0" shows "convergent (\N. integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n))" (is "convergent (?f n)") proof - have "convergent (?f (Suc n))" for n proof (induction n) case 0 thus ?case using integral_pbernpoly_1 by (auto intro!: convergentI) next case (Suc n) have "convergent (\N. ?f (Suc n) N - of_real (pbernpoly (Suc (Suc n)) (real N)) / (of_nat (Suc (Suc n)) * (s + of_nat N) ^ Suc n) + of_real (bernoulli (Suc (Suc n)) / (real (Suc (Suc n)))) / s ^ Suc n)" (is "convergent ?g") by (intro convergent_add convergent_diff Suc convergent_const convergentI[OF pbernpoly_over_power_tendsto_0]) simp_all also have "?g = (\N. of_nat (Suc n) / of_nat (Suc (Suc n)) * ?f (Suc (Suc n)) N)" using s by (subst pbernpoly_integral_conv_pbernpoly_integral_Suc) (auto simp: fun_eq_iff field_simps simp del: of_nat_Suc power_Suc) also have "convergent \ \ convergent (?f (Suc (Suc n)))" by (intro convergent_mult_const_iff) (simp_all del: of_nat_Suc) finally show ?case . qed from this[of "n - 1"] assms show ?thesis by simp qed lemma stirling_integral_conv_stirling_integral_Suc: assumes "n > 0" shows "stirling_integral n s = of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s - of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n)" proof - have "(\N. of_real (pbernpoly (Suc n) (real N)) / (of_nat (Suc n) * (s + of_nat N) ^ n) - of_real (bernoulli (Suc n)) / (real (Suc n) * s ^ n) + integral {0..real N} (\x. of_nat n / of_nat (Suc n) * (of_real (pbernpoly (Suc n) x) / (of_real x + s) ^ Suc n))) \ 0 - of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n) + of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s" (is "?f \ _") unfolding stirling_integral_def integral_mult_right using convergent_stirling_integral[of "Suc n"] assms s by (intro tendsto_intros pbernpoly_over_power_tendsto_0) (auto simp: convergent_LIMSEQ_iff simp del: of_nat_Suc) also have "?this \ (\N. integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s - of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n)" using eventually_gt_at_top[of "0::nat"] pbernpoly_integral_conv_pbernpoly_integral_Suc[of n] assms unfolding integral_mult_right by (intro filterlim_cong refl) (auto elim!: eventually_mono simp del: power_Suc) finally show ?thesis unfolding stirling_integral_def[of n] by (rule limI) qed lemma stirling_integral_1_unfold: assumes "m > 0" shows "stirling_integral 1 s = stirling_integral m s / of_nat m - (\k=1..k=1..k = 1.. 0" shows "ln_Gamma s = (s - 1 / 2) * ln s - s + ln (2 * pi) / 2 + (\k=1..k = 1.. 0 \ (\x. integral {0..real x} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ stirling_integral n s" unfolding stirling_integral_def using convergent_stirling_integral[of n] by (simp only: convergent_LIMSEQ_iff) end lemmas has_integral_of_real = has_integral_linear[OF _ bounded_linear_of_real, unfolded o_def] lemmas integral_of_real = integral_linear[OF _ bounded_linear_of_real, unfolded o_def] lemma integrable_ln_Gamma_aux_real: assumes "0 < s" shows "(\x. pbernpoly n x / (x + s) ^ n) integrable_on {0..real N}" proof - have "(\x. complex_of_real (pbernpoly n x / (x + s) ^ n)) integrable_on {0..real N}" using integrable_ln_Gamma_aux[of "of_real s" n N] assms by simp from integrable_linear[OF this bounded_linear_Re] show ?thesis by (simp only: o_def Re_complex_of_real) qed lemma assumes "x > 0" "n > 0" shows stirling_integral_complex_of_real: "stirling_integral n (complex_of_real x) = of_real (stirling_integral n x)" and LIMSEQ_stirling_integral_real: "(\N. integral {0..real N} (\t. pbernpoly n t / (t + x) ^ n)) \ stirling_integral n x" and stirling_integral_real_convergent: "convergent (\N. integral {0..real N} (\t. pbernpoly n t / (t + x) ^ n))" proof - have "(\N. integral {0..real N} (\t. of_real (pbernpoly n t / (t + x) ^ n))) \ stirling_integral n (complex_of_real x)" using LIMSEQ_stirling_integral[of "complex_of_real x" n] assms by simp hence "(\N. of_real (integral {0..real N} (\t. pbernpoly n t / (t + x) ^ n))) \ stirling_integral n (complex_of_real x)" using integrable_ln_Gamma_aux_real[OF assms(1), of n] by (subst (asm) integral_of_real) simp from tendsto_Re[OF this] have "(\N. integral {0..real N} (\t. pbernpoly n t / (t + x) ^ n)) \ Re (stirling_integral n (complex_of_real x))" by simp thus "convergent (\N. integral {0..real N} (\t. pbernpoly n t / (t + x) ^ n))" by (rule convergentI) thus "(\N. integral {0..real N} (\t. pbernpoly n t / (t + x) ^ n)) \ stirling_integral n x" unfolding stirling_integral_def by (simp add: convergent_LIMSEQ_iff) from tendsto_of_real[OF this, where 'a = complex] integrable_ln_Gamma_aux_real[OF assms(1), of n] have "(\xa. integral {0..real xa} (\xa. complex_of_real (pbernpoly n xa) / (complex_of_real xa + x) ^ n)) \ complex_of_real (stirling_integral n x)" by (subst (asm) integral_of_real [symmetric]) simp_all from LIMSEQ_unique[OF this LIMSEQ_stirling_integral[of "complex_of_real x" n]] assms show "stirling_integral n (complex_of_real x) = of_real (stirling_integral n x)" by simp qed lemma ln_Gamma_stirling_real: assumes "x > (0 :: real)" "m > (0::nat)" shows "ln_Gamma x = (x - 1 / 2) * ln x - x + ln (2 * pi) / 2 + (\k = 1..k = 1.. (1::nat)" obtains c where "\s. Re s > 0 \ norm (stirling_integral n s) \ c / Re s ^ (n - 1)" proof - obtain c where c: "norm (pbernpoly n x) \ c" for x by (rule bounded_pbernpoly[of n]) blast have c': "pbernpoly n x \ c" for x using c[of x] by (simp add: abs_real_def split: if_splits) from c[of 0] have c_nonneg: "c \ 0" by simp have "norm (stirling_integral n s) \ c / (real n - 1) / Re s ^ (n - 1)" if s: "Re s > 0" for s proof (rule Lim_norm_ubound[OF _ LIMSEQ_stirling_integral]) have pos: "x + norm s > 0" if "x \ 0" for x using s that by (intro add_nonneg_pos) auto have nz: "of_real x + s \ 0" if "x \ 0" for x using s that by (auto simp: complex_eq_iff) let ?bound = "\N. c / (Re s ^ (n - 1) * (real n - 1)) - c / ((real N + Re s) ^ (n - 1) * (real n - 1))" show "eventually (\N. norm (integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ c / (real n - 1) / Re s ^ (n - 1)) at_top" using eventually_gt_at_top[of "0::nat"] proof eventually_elim case (elim N) let ?F = "\x. -c / ((x + Re s) ^ (n - 1) * (real n - 1))" from n s have "((\x. c / (x + Re s) ^ n) has_integral (?F (real N) - ?F 0)) {0..real N}" by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros simp: divide_simps power_diff add_eq_0_iff2 has_field_derivative_iff_has_vector_derivative [symmetric]) also have "?F (real N) - ?F 0 = ?bound N" by simp finally have *: "((\x. c / (x + Re s) ^ n) has_integral ?bound N) {0..real N}" . have "norm (integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ integral {0..real N} (\x. c / (x + Re s) ^ n)" proof (intro integral_norm_bound_integral integrable_ln_Gamma_aux s ballI) fix x assume x: "x \ {0..real N}" have "norm (of_real (pbernpoly n x) / (of_real x + s) ^ n) \ c / norm (of_real x + s) ^ n" unfolding norm_divide norm_power using c by (intro divide_right_mono) simp_all also have "\ \ c / (x + Re s) ^ n" using x c c_nonneg s nz[of x] complex_Re_le_cmod[of "of_real x + s"] by (intro divide_left_mono power_mono mult_pos_pos zero_less_power add_nonneg_pos) auto finally show "norm (of_real (pbernpoly n x) / (of_real x + s) ^ n) \ \" . qed (insert n s * pos nz c, auto simp: complex_nonpos_Reals_iff) also have "\ = ?bound N" using * by (simp add: has_integral_iff) also have "\ \ c / (Re s ^ (n - 1) * (real n - 1))" using c_nonneg elim s n by simp also have "\ = c / (real n - 1) / (Re s ^ (n - 1))" by simp finally show "norm (integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ c / (real n - 1) / Re s ^ (n - 1)" . qed qed (insert s n, simp_all add: complex_nonpos_Reals_iff) thus ?thesis by (rule that) qed lemma stirling_integral_bound_aux_integral1: fixes a b c :: real and n :: nat assumes "a \ 0" "b > 0" "c \ 0" "n > 1" "l < a - b" "r > a + b" shows "((\x. c / max b \x - a\ ^ n) has_integral 2*c*(n / (n - 1))/b^(n-1) - c/(n-1) * (1/(a-l)^(n-1) + 1/(r-a)^(n-1))) {l..r}" proof - define x1 x2 where "x1 = a - b" and "x2 = a + b" define F1 where "F1 = (\x::real. c / (a - x) ^ (n - 1) / (n - 1))" define F3 where "F3 = (\x::real. -c / (x - a) ^ (n - 1) / (n - 1))" have deriv: "(F1 has_vector_derivative (c / (a - x) ^ n)) (at x within A)" "(F3 has_vector_derivative (c / (x - a) ^ n)) (at x within A)" if "x \ a" for x :: real and A unfolding F1_def F3_def using assms that by (auto intro!: derivative_eq_intros simp: divide_simps power_diff add_eq_0_iff2 simp flip: has_field_derivative_iff_has_vector_derivative) from assms have "((\x. c / (a - x) ^ n) has_integral (F1 x1 - F1 l)) {l..x1}" by (intro fundamental_theorem_of_calculus deriv) (auto simp: x1_def max_def split: if_splits) also have "?this \ ((\x. c / max b \x - a\ ^ n) has_integral (F1 x1 - F1 l)) {l..x1}" using assms by (intro has_integral_spike_finite_eq[of "{l}"]) (auto simp: x1_def max_def split: if_splits) finally have I1: "((\x. c / max b \x - a\ ^ n) has_integral (F1 x1 - F1 l)) {l..x1}" . have "((\x. c / b ^ n) has_integral (x2 - x1) * c / b ^ n) {x1..x2}" using has_integral_const_real[of "c / b ^ n" x1 x2] assms by (simp add: x1_def x2_def) also have "?this \ ((\x. c / max b \x - a\ ^ n) has_integral ((x2 - x1) * c / b ^ n)) {x1..x2}" by (intro has_integral_spike_finite_eq[of "{x1, x2}"]) (auto simp: x1_def x2_def split: if_splits) finally have I2: "((\x. c / max b \x - a\ ^ n) has_integral ((x2 - x1) * c / b ^ n)) {x1..x2}" . from assms have I3: "((\x. c / (x - a) ^ n) has_integral (F3 r - F3 x2)) {x2..r}" by (intro fundamental_theorem_of_calculus deriv) (auto simp: x2_def min_def split: if_splits) also have "?this \ ((\x. c / max b \x - a\ ^ n) has_integral (F3 r - F3 x2)) {x2..r}" using assms by (intro has_integral_spike_finite_eq[of "{r}"]) (auto simp: x2_def min_def split: if_splits) finally have I3: "((\x. c / max b \x - a\ ^ n) has_integral (F3 r - F3 x2)) {x2..r}" . have "((\x. c / max b \x - a\ ^ n) has_integral (F1 x1 - F1 l) + ((x2 - x1) * c / b ^ n) + (F3 r - F3 x2)) {l..r}" using assms by (intro has_integral_combine[OF _ _ has_integral_combine[OF _ _ I1 I2] I3]) (auto simp: x1_def x2_def) also have "(F1 x1 - F1 l) + ((x2 - x1) * c / b ^ n) + (F3 r - F3 x2) = F1 x1 - F1 l + F3 r - F3 x2 + (x2 - x1) * c / b ^ n" by (simp add: algebra_simps) also have "x2 - x1 = 2 * b" using assms by (simp add: x2_def x1_def min_def max_def) also have "2 * b * c / b ^ n = 2 * c / b ^ (n - 1)" using assms by (simp add: power_diff field_simps) also have "F1 x1 - F1 l + F3 r - F3 x2 = c/(n-1) * (2/b^(n-1) - 1/(a-l)^(n-1) - 1/(r-a)^(n-1))" using assms by (simp add: x1_def x2_def F1_def F3_def field_simps) also have "\ + 2 * c / b ^ (n - 1) = 2*c*(1 + 1/(n-1))/b^(n-1) - c/(n-1) * (1/(a-l)^(n-1) + 1/(r-a)^(n-1))" using assms by (simp add: field_simps) also have "1 + 1 / (n - 1) = n / (n - 1)" using assms by (simp add: field_simps) finally show ?thesis . qed lemma stirling_integral_bound_aux_integral2: fixes a b c :: real and n :: nat assumes "a \ 0" "b > 0" "c \ 0" "n > 1" obtains I where "((\x. c / max b \x - a\ ^ n) has_integral I) {l..r}" "I \ 2 * c * (n / (n - 1)) / b ^ (n-1)" proof - define l' where "l' = min l (a - b - 1)" define r' where "r' = max r (a + b + 1)" define A where "A = 2 * c * (n / (n - 1)) / b ^ (n - 1)" define B where "B = c / real (n - 1) * (1 / (a - l') ^ (n - 1) + 1 / (r' - a) ^ (n - 1))" have has_int: "((\x. c / max b \x - a\ ^ n) has_integral (A - B)) {l'..r'}" using assms unfolding A_def B_def by (intro stirling_integral_bound_aux_integral1) (auto simp: l'_def r'_def) have "(\x. c / max b \x - a\ ^ n) integrable_on {l..r}" by (rule integrable_on_subinterval[OF has_integral_integrable[OF has_int]]) (auto simp: l'_def r'_def) then obtain I where has_int': "((\x. c / max b \x - a\ ^ n) has_integral I) {l..r}" by (auto simp: integrable_on_def) from assms have "I \ A - B" by (intro has_integral_subset_le[OF _ has_int' has_int]) (auto simp: l'_def r'_def) also have "\ \ A" using assms by (simp add: B_def l'_def r'_def) finally show ?thesis using that[of I] has_int' unfolding A_def by blast qed lemma stirling_integral_bound_aux': assumes n: "n > (1::nat)" and \: "\ \ {0<..s::complex. s \ complex_cone' \ - {0} \ norm (stirling_integral n s) \ c / norm s ^ (n - 1)" proof - obtain c where c: "norm (pbernpoly n x) \ c" for x by (rule bounded_pbernpoly[of n]) blast have c': "pbernpoly n x \ c" for x using c[of x] by (simp add: abs_real_def split: if_splits) from c[of 0] have c_nonneg: "c \ 0" by simp define D where "D = c * Beta (- (real_of_int (- int n) / 2) - 1 / 2) (1 / 2) / 2" define C where "C = max D (2*c*(n/(n-1))/sin \^(n-1))" have *: "norm (stirling_integral n s) \ C / norm s ^ (n - 1)" if s: "s \ complex_cone' \ - {0}" for s :: complex proof (rule Lim_norm_ubound[OF _ LIMSEQ_stirling_integral]) - from s \ have arg: "\arg s\ \ \" by (auto simp: complex_cone_altdef) + from s \ have Arg: "\Arg s\ \ \" by (auto simp: complex_cone_altdef) have s': "s \ \\<^sub>\\<^sub>0" using complex_cone_inter_nonpos_Reals[of "-\" \] \ s by auto from s have [simp]: "s \ 0" by auto show "eventually (\N. norm (integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ C / norm s ^ (n - 1)) at_top" using eventually_gt_at_top[of "0::nat"] proof eventually_elim case (elim N) show ?case proof (cases "Re s > 0") case True have int: "((\x. c * (x^2 + norm s^2) powr (-n / 2)) has_integral D * (norm s ^ 2) powr (-n / 2 + 1 / 2)) {0<..}" using has_integral_mult_left[OF has_integral_Beta3[of "-n/2" "norm s ^ 2"], of c] assms True unfolding D_def by (simp add: algebra_simps) hence int': "((\x. c * (x^2 + norm s^2) powr (-n / 2)) has_integral D * (norm s ^ 2) powr (-n / 2 + 1 / 2)) {0..}" by (subst has_integral_interior [symmetric]) simp_all hence integrable: "(\x. c * (x^2 + norm s^2) powr (-n / 2)) integrable_on {0..}" by (simp add: has_integral_iff) have "norm (integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ integral {0..real N} (\x. c * (x^2 + norm s^2) powr (-n / 2))" proof (intro integral_norm_bound_integral s ballI integrable_ln_Gamma_aux) have [simp]: "{0<..} - {0::real..} = {}" "{0..} - {0<..} = {0::real}" by auto have "(\x. c * (x\<^sup>2 + (cmod s)\<^sup>2) powr (real_of_int (- int n) / 2)) integrable_on {0<..}" using int by (simp add: has_integral_iff) also have "?this \ (\x. c * (x\<^sup>2 + (cmod s)\<^sup>2) powr (real_of_int (- int n) / 2)) integrable_on {0..}" by (intro integrable_spike_set_eq) auto finally show "(\x. c * (x\<^sup>2 + (cmod s)\<^sup>2) powr (real_of_int (- int n) / 2)) integrable_on {0..real N}" by (rule integrable_on_subinterval) auto next fix x assume x: "x \ {0..real N}" have nz: "complex_of_real x + s \ 0" using True x by (auto simp: complex_eq_iff) have "norm (of_real (pbernpoly n x) / (of_real x + s) ^ n) \ c / norm (of_real x + s) ^ n" unfolding norm_divide norm_power using c by (intro divide_right_mono) simp_all also have "\ \ c / sqrt (x ^ 2 + norm s ^ 2) ^ n" proof (intro divide_left_mono mult_pos_pos zero_less_power power_mono) show "sqrt (x\<^sup>2 + (cmod s)\<^sup>2) \ cmod (complex_of_real x + s)" using x True by (simp add: cmod_def algebra_simps power2_eq_square) qed (use x True c_nonneg assms nz in \auto simp: add_nonneg_pos\) also have "sqrt (x ^ 2 + norm s ^ 2) ^ n = (x ^ 2 + norm s ^ 2) powr (1/2 * n)" by (subst powr_powr [symmetric], subst powr_realpow) (auto simp: powr_half_sqrt add_nonneg_pos) also have "c / \ = c * (x^2 + norm s^2) powr (-n / 2)" by (simp add: powr_minus field_simps) finally show "norm (complex_of_real (pbernpoly n x) / (complex_of_real x + s) ^ n) \ \" . qed fact+ also have "\ \ integral {0..} (\x. c * (x^2 + norm s^2) powr (-n / 2))" using c_nonneg by (intro integral_subset_le integrable integrable_on_subinterval[OF integrable]) auto also have "\ = D * (norm s ^ 2) powr (-n / 2 + 1 / 2)" using int' by (simp add: has_integral_iff) also have "(norm s ^ 2) powr (-n / 2 + 1 / 2) = norm s powr (2 * (-n / 2 + 1 / 2))" by (subst powr_powr [symmetric]) auto also have "\ = norm s powr (-real (n - 1))" using assms by (simp add: of_nat_diff) also have "D * \ = D / norm s ^ (n - 1)" by (auto simp: powr_minus powr_realpow field_simps) also have "\ \ C / norm s ^ (n - 1)" by (intro divide_right_mono) (auto simp: C_def) finally show "norm (integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ \" . next case False - have "cos \arg s\ = cos (arg s)" + have "cos \Arg s\ = cos (Arg s)" by (simp add: abs_if) - also have "cos (arg s) = Re (rcis (norm s) (arg s)) / norm s" + also have "cos (Arg s) = Re (rcis (norm s) (Arg s)) / norm s" by (subst Re_rcis) auto also have "\ = Re s / norm s" - by (subst rcis_cmod_arg) auto + by (subst rcis_cmod_Arg) auto also have "\ \ cos (pi / 2)" using False by (auto simp: field_simps) - finally have "\arg s\ \ pi / 2" - using arg \ by (subst (asm) cos_mono_le_eq) auto + finally have "\Arg s\ \ pi / 2" + using Arg \ by (subst (asm) cos_mono_le_eq) auto have "sin \ * norm s = sin (pi - \) * norm s" by simp - also have "\ \ sin (pi - \arg s\) * norm s" - using \ arg \\arg s\ \ pi / 2\ + also have "\ \ sin (pi - \Arg s\) * norm s" + using \ Arg \\Arg s\ \ pi / 2\ by (intro mult_right_mono sin_monotone_2pi_le) auto - also have "sin \arg s\ \ 0" - using arg_bounded[of s] by (intro sin_ge_zero) auto - hence "sin (pi - \arg s\) = \sin \arg s\\" + also have "sin \Arg s\ \ 0" + using Arg_bounded[of s] by (intro sin_ge_zero) auto + hence "sin (pi - \Arg s\) = \sin \Arg s\\" by simp - also have "\ = \sin (arg s)\" + also have "\ = \sin (Arg s)\" by (simp add: abs_if) - also have "\ * norm s = \Im (rcis (norm s) (arg s))\" + also have "\ * norm s = \Im (rcis (norm s) (Arg s))\" by (simp add: abs_mult) also have "\ = \Im s\" - by (subst rcis_cmod_arg) auto + by (subst rcis_cmod_Arg) auto finally have abs_Im_ge: "\Im s\ \ sin \ * norm s" . have [simp]: "Im s \ 0" "s \ 0" using s \s \ \\<^sub>\\<^sub>0\ False by (auto simp: cmod_def zero_le_mult_iff complex_nonpos_Reals_iff) have "sin \ > 0" using assms by (intro sin_gt_zero) auto obtain I where I: "((\x. c / max \Im s\ \x + Re s\ ^ n) has_integral I) {0..real N}" "I \ 2*c*(n/(n-1)) / \Im s\ ^ (n - 1)" using s c_nonneg assms False stirling_integral_bound_aux_integral2[of "-Re s" "\Im s\" c n 0 "real N"] by auto have "norm (integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ integral {0..real N} (\x. c / max \Im s\ \x + Re s\ ^ n)" proof (intro integral_norm_bound_integral integrable_ln_Gamma_aux s ballI) show "(\x. c / max \Im s\ \x + Re s\ ^ n) integrable_on {0..real N}" using I(1) by (simp add: has_integral_iff) next fix x assume x: "x \ {0..real N}" have nz: "complex_of_real x + s \ 0" by (auto simp: complex_eq_iff) have "norm (complex_of_real (pbernpoly n x) / (complex_of_real x + s) ^ n) \ c / norm (complex_of_real x + s) ^ n" unfolding norm_divide norm_power using c[of x] by (intro divide_right_mono) simp_all also have "\ \ c / max \Im s\ \x + Re s\ ^ n" using c_nonneg nz abs_Re_le_cmod[of "of_real x + s"] abs_Im_le_cmod[of "of_real x + s"] by (intro divide_left_mono power_mono mult_pos_pos zero_less_power) (auto simp: less_max_iff_disj) finally show "norm (complex_of_real (pbernpoly n x) / (complex_of_real x + s) ^ n) \ \" . qed (auto simp: complex_nonpos_Reals_iff) also have "\ \ 2*c*(n/(n-1)) / \Im s\ ^ (n - 1)" using I by (simp add: has_integral_iff) also have "\ \ 2*c*(n/(n-1)) / (sin \ * norm s) ^ (n - 1)" using \sin \ > 0\ s c_nonneg abs_Im_ge by (intro divide_left_mono mult_pos_pos zero_less_power power_mono mult_nonneg_nonneg) auto also have "\ = 2*c*(n/(n-1))/sin \^(n-1) / norm s ^ (n - 1)" by (simp add: field_simps) also have "\ \ C / norm s ^ (n - 1)" by (intro divide_right_mono) (auto simp: C_def) finally show ?thesis . qed qed qed (use that assms complex_cone_inter_nonpos_Reals[of "-\" \] \ in auto) thus ?thesis by (rule that) qed lemma stirling_integral_bound: assumes "n > 0" obtains c where "\s. Re s > 0 \ norm (stirling_integral n s) \ c / Re s ^ n" proof - let ?f = "\s. of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s - of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n)" from stirling_integral_bound_aux[of "Suc n"] assms obtain c where c: "\s. Re s > 0 \ norm (stirling_integral (Suc n) s) \ c / Re s ^ n" by auto define c1 where "c1 = real n / real (Suc n) * c" define c2 where "c2 = \bernoulli (Suc n)\ / real (Suc n)" have c2_nonneg: "c2 \ 0" by (simp add: c2_def) show ?thesis proof (rule that) fix s :: complex assume s: "Re s > 0" hence s': "s \ \\<^sub>\\<^sub>0" by (auto simp: complex_nonpos_Reals_iff) have "stirling_integral n s = ?f s" using s' assms by (rule stirling_integral_conv_stirling_integral_Suc) also have "norm \ \ norm (of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s) + norm (of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n))" by (rule norm_triangle_ineq4) also have "\ = real n / real (Suc n) * norm (stirling_integral (Suc n) s) + c2 / norm s ^ n" (is "_ = ?A + ?B") by (simp add: norm_divide norm_mult norm_power c2_def field_simps del: of_nat_Suc) also have "?A \ real n / real (Suc n) * (c / Re s ^ n)" by (intro mult_left_mono c s) simp_all also have "\ = c1 / Re s ^ n" by (simp add: c1_def) also have "c2 / norm s ^ n \ c2 / Re s ^ n" using s c2_nonneg by (intro divide_left_mono power_mono complex_Re_le_cmod mult_pos_pos zero_less_power) auto also have "c1 / Re s ^ n + c2 / Re s ^ n = (c1 + c2) / Re s ^ n" using s by (simp add: field_simps) finally show "norm (stirling_integral n s) \ (c1 + c2) / Re s ^ n" by - simp_all qed qed lemma stirling_integral_bound': assumes "n > 0" and "\ \ {0<..s::complex. s \ complex_cone' \ - {0} \ norm (stirling_integral n s) \ c / norm s ^ n" proof - let ?f = "\s. of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s - of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n)" from stirling_integral_bound_aux'[of "Suc n"] assms obtain c where c: "\s::complex. s \ complex_cone' \ - {0} \ norm (stirling_integral (Suc n) s) \ c / norm s ^ n" by auto define c1 where "c1 = real n / real (Suc n) * c" define c2 where "c2 = \bernoulli (Suc n)\ / real (Suc n)" have c2_nonneg: "c2 \ 0" by (simp add: c2_def) show ?thesis proof (rule that) fix s :: complex assume s: "s \ complex_cone' \ - {0}" have s': "s \ \\<^sub>\\<^sub>0" using complex_cone_inter_nonpos_Reals[of "-\" \] assms s by auto have "stirling_integral n s = ?f s" using s' assms by (intro stirling_integral_conv_stirling_integral_Suc) auto also have "norm \ \ norm (of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s) + norm (of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n))" by (rule norm_triangle_ineq4) also have "\ = real n / real (Suc n) * norm (stirling_integral (Suc n) s) + c2 / norm s ^ n" (is "_ = ?A + ?B") by (simp add: norm_divide norm_mult norm_power c2_def field_simps del: of_nat_Suc) also have "?A \ real n / real (Suc n) * (c / norm s ^ n)" by (intro mult_left_mono c s) simp_all also have "\ = c1 / norm s ^ n" by (simp add: c1_def) also have "c1 / norm s ^ n + c2 / norm s ^ n = (c1 + c2) / norm s ^ n" using s by (simp add: divide_simps) finally show "norm (stirling_integral n s) \ (c1 + c2) / norm s ^ n" by - simp_all qed qed lemma stirling_integral_holomorphic [holomorphic_intros]: assumes m: "m > 0" and "A \ \\<^sub>\\<^sub>0 = {}" shows "stirling_integral m holomorphic_on A" proof - from assms have [simp]: "z \ \\<^sub>\\<^sub>0" if "z \ A" for z using that by auto let ?f = "\s::complex. of_nat m * ((s - 1 / 2) * Ln s - s + of_real (ln (2 * pi) / 2) + (\k=1.. stirling_integral m holomorphic_on A" using assms by (intro holomorphic_cong refl) (simp_all add: field_simps ln_Gamma_stirling_complex) finally show "stirling_integral m holomorphic_on A" . qed lemma stirling_integral_continuous_on_complex [continuous_intros]: assumes m: "m > 0" and "A \ \\<^sub>\\<^sub>0 = {}" shows "continuous_on A (stirling_integral m :: _ \ complex)" by (intro holomorphic_on_imp_continuous_on stirling_integral_holomorphic assms) lemma has_field_derivative_stirling_integral_complex: fixes x :: complex assumes "x \ \\<^sub>\\<^sub>0" "n > 0" shows "(stirling_integral n has_field_derivative deriv (stirling_integral n) x) (at x)" using assms by (intro holomorphic_derivI[OF stirling_integral_holomorphic, of n "-\\<^sub>\\<^sub>0"]) auto lemma assumes n: "n > 0" and "x > 0" shows deriv_stirling_integral_complex_of_real: "(deriv ^^ j) (stirling_integral n) (complex_of_real x) = complex_of_real ((deriv ^^ j) (stirling_integral n) x)" (is "?lhs x = ?rhs x") and differentiable_stirling_integral_real: "(deriv ^^ j) (stirling_integral n) field_differentiable at x" (is ?thesis2) proof - let ?A = "{s. Re s > 0}" let ?f = "\j x. (deriv ^^ j) (stirling_integral n) (complex_of_real x)" let ?f' = "\j x. complex_of_real ((deriv ^^ j) (stirling_integral n) x)" have [simp]: "open ?A" by (simp add: open_halfspace_Re_gt) have "?lhs x = ?rhs x \ (deriv ^^ j) (stirling_integral n) field_differentiable at x" if "x > 0" for x using that proof (induction j arbitrary: x) case 0 have "((\x. Re (stirling_integral n (of_real x))) has_field_derivative Re (deriv (\x. stirling_integral n x) (of_real x))) (at x)" using 0 n by (auto intro!: derivative_intros has_vector_derivative_real_field field_differentiable_derivI holomorphic_on_imp_differentiable_at[of _ ?A] stirling_integral_holomorphic simp: complex_nonpos_Reals_iff) also have "?this \ (stirling_integral n has_field_derivative Re (deriv (\x. stirling_integral n x) (of_real x))) (at x)" using eventually_nhds_in_open[of "{0<..}" x] 0 n by (intro has_field_derivative_cong_ev refl) (auto elim!: eventually_mono simp: stirling_integral_complex_of_real) finally have "stirling_integral n field_differentiable at x" by (auto simp: field_differentiable_def) with 0 n show ?case by (auto simp: stirling_integral_complex_of_real) next case (Suc j x) note IH = conjunct1[OF Suc.IH] conjunct2[OF Suc.IH] have *: "(deriv ^^ Suc j) (stirling_integral n) (complex_of_real x) = of_real ((deriv ^^ Suc j) (stirling_integral n) x)" if x: "x > 0" for x proof - have "deriv ((deriv ^^ j) (stirling_integral n)) (complex_of_real x) = vector_derivative (\x. (deriv ^^ j) (stirling_integral n) (of_real x)) (at x)" using n x by (intro vector_derivative_of_real_right [symmetric] holomorphic_on_imp_differentiable_at[of _ ?A] holomorphic_higher_deriv stirling_integral_holomorphic) (auto simp: complex_nonpos_Reals_iff) also have "\ = vector_derivative (\x. of_real ((deriv ^^ j) (stirling_integral n) x)) (at x)" using eventually_nhds_in_open[of "{0<..}" x] x by (intro vector_derivative_cong_eq) (auto elim!: eventually_mono simp: IH(1)) also have "\ = of_real (deriv ((deriv ^^ j) (stirling_integral n)) x)" by (intro vector_derivative_of_real_left holomorphic_on_imp_differentiable_at[of _ ?A] field_differentiable_imp_differentiable IH(2) x) finally show ?thesis by simp qed have "((\x. Re ((deriv ^^ Suc j) (stirling_integral n) (of_real x))) has_field_derivative Re (deriv ((deriv ^^ Suc j) (stirling_integral n)) (of_real x))) (at x)" using Suc.prems n by (intro derivative_intros has_vector_derivative_real_field field_differentiable_derivI holomorphic_on_imp_differentiable_at[of _ ?A] stirling_integral_holomorphic holomorphic_higher_deriv) (auto simp: complex_nonpos_Reals_iff) also have "?this \ ((deriv ^^ Suc j) (stirling_integral n) has_field_derivative Re (deriv ((deriv ^^ Suc j) (stirling_integral n)) (of_real x))) (at x)" using eventually_nhds_in_open[of "{0<..}" x] Suc.prems * by (intro has_field_derivative_cong_ev refl) (auto elim!: eventually_mono) finally have "(deriv ^^ Suc j) (stirling_integral n) field_differentiable at x" by (auto simp: field_differentiable_def) with *[OF Suc.prems] show ?case by blast qed from this[OF assms(2)] show "?lhs x = ?rhs x" ?thesis2 by blast+ qed text \ Unfortunately, asymptotic power series cannot, in general, be differentiated. However, since @{term ln_Gamma} is holomorphic on the entire positive real half-space, we can differentiate its asymptotic expansion after all. To do this, we use an ad-hoc version of the more general approach outlined in Erdelyi's ``Asymptotic Expansions'' for holomorphic functions: We bound the value of the $j$-th derivative of the remainder term at some value $x$ by applying Cauchy's integral formula along a circle centred at $x$ with radius $\frac{1}{2} x$. \ lemma deriv_stirling_integral_real_bound: assumes m: "m > 0" shows "(deriv ^^ j) (stirling_integral m) \ O(\x::real. 1 / x ^ (m + j))" proof - from stirling_integral_bound[OF m] guess c . note c = this have "0 \ cmod (stirling_integral m 1)" by simp also have "\ \ c" using c[of 1] by simp finally have c_nonneg: "c \ 0" . define B where "B = c * 2 ^ (m + Suc j)" define B' where "B' = B * fact j / 2" have "eventually (\x::real. norm ((deriv ^^ j) (stirling_integral m) x) \ B' * norm (1 / x ^ (m+ j))) at_top" using eventually_gt_at_top[of "0::real"] proof eventually_elim case (elim x) have "s \ \\<^sub>\\<^sub>0" if "s \ cball (of_real x) (x/2)" for s :: complex proof - have "x - Re s \ norm (of_real x - s)" using complex_Re_le_cmod[of "of_real x - s"] by simp also from that have "\ \ x/2" by (simp add: dist_complex_def) finally show ?thesis using elim by (auto simp: complex_nonpos_Reals_iff) qed hence "((\u. stirling_integral m u / (u - of_real x) ^ Suc j) has_contour_integral complex_of_real (2 * pi) * \ / fact j * (deriv ^^ j) (stirling_integral m) (of_real x)) (circlepath (of_real x) (x/2))" using m elim by (intro Cauchy_has_contour_integral_higher_derivative_circlepath stirling_integral_continuous_on_complex stirling_integral_holomorphic) auto hence "norm (of_real (2 * pi) * \ / fact j * (deriv ^^ j) (stirling_integral m) (of_real x)) \ B / x ^ (m + Suc j) * (2 * pi * (x / 2))" proof (rule has_contour_integral_bound_circlepath) fix u :: complex assume dist: "norm (u - of_real x) = x / 2" have "Re (of_real x - u) \ norm (of_real x - u)" by (rule complex_Re_le_cmod) also have "\ = x / 2" using dist by (simp add: norm_minus_commute) finally have Re_u: "Re u \ x/2" using elim by simp have "norm (stirling_integral m u / (u - of_real x) ^ Suc j) \ c / Re u ^ m / (x / 2) ^ Suc j" using Re_u elim unfolding norm_divide norm_power dist by (intro divide_right_mono zero_le_power c) simp_all also have "\ \ c / (x/2) ^ m / (x / 2) ^ Suc j" using c_nonneg elim Re_u by (intro divide_right_mono divide_left_mono power_mono) simp_all also have "\ = B / x ^ (m + Suc j)" using elim by (simp add: B_def field_simps power_add) finally show "norm (stirling_integral m u / (u - of_real x) ^ Suc j) \ B / x ^ (m + Suc j)" . qed (insert elim c_nonneg, auto simp: B_def simp del: power_Suc) hence "cmod ((deriv ^^ j) (stirling_integral m) (of_real x)) \ B' / x ^ (j + m)" using elim by (simp add: field_simps norm_divide norm_mult norm_power B'_def) with elim m show ?case by (simp_all add: add_ac deriv_stirling_integral_complex_of_real) qed thus ?thesis by (rule bigoI) qed definition stirling_sum where "stirling_sum j m x = (-1) ^ j * (\k = 1..k\m. (of_real (bernoulli' k) * pochhammer (of_nat (Suc k)) (j - 1) * inverse x ^ (k + j)))" lemma stirling_sum_complex_of_real: "stirling_sum j m (complex_of_real x) = complex_of_real (stirling_sum j m x)" by (simp add: stirling_sum_def pochhammer_of_real [symmetric] del: of_nat_Suc) lemma stirling_sum'_complex_of_real: "stirling_sum' j m (complex_of_real x) = complex_of_real (stirling_sum' j m x)" by (simp add: stirling_sum'_def pochhammer_of_real [symmetric] del: of_nat_Suc) lemma has_field_derivative_stirling_sum_complex [derivative_intros]: "Re x > 0 \ (stirling_sum j m has_field_derivative stirling_sum (Suc j) m x) (at x)" unfolding stirling_sum_def [abs_def] sum_distrib_left by (rule DERIV_sum) (auto intro!: derivative_eq_intros simp del: of_nat_Suc simp: pochhammer_Suc power_diff) lemma has_field_derivative_stirling_sum_real [derivative_intros]: "x > (0::real) \ (stirling_sum j m has_field_derivative stirling_sum (Suc j) m x) (at x)" unfolding stirling_sum_def [abs_def] sum_distrib_left by (rule DERIV_sum) (auto intro!: derivative_eq_intros simp del: of_nat_Suc simp: pochhammer_Suc power_diff) lemma has_field_derivative_stirling_sum'_complex [derivative_intros]: assumes "j > 0" "Re x > 0" shows "(stirling_sum' j m has_field_derivative stirling_sum' (Suc j) m x) (at x)" proof (cases j) case (Suc j') from assms have [simp]: "x \ 0" by auto define c where "c = (\n. (-1) ^ Suc j * complex_of_real (bernoulli' n) * pochhammer (of_nat (Suc n)) j')" define T where "T = (\n x. c n * inverse x ^ (j + n))" define T' where "T' = (\n x. - (of_nat (j + n)) * c n * inverse x ^ (Suc (j + n)))" have "((\x. \k\m. T k x) has_field_derivative (\k\m. T' k x)) (at x)" using assms Suc by (intro DERIV_sum) (auto simp: T_def T'_def intro!: derivative_eq_intros simp: field_simps power_add [symmetric] simp del: of_nat_Suc power_Suc of_nat_add) also have "(\x. (\k\m. T k x)) = stirling_sum' j m" by (simp add: Suc T_def c_def stirling_sum'_def fun_eq_iff add_ac mult.assoc sum_distrib_left) also have "(\k\m. T' k x) = stirling_sum' (Suc j) m x" by (simp add: T'_def c_def Suc stirling_sum'_def sum_distrib_left sum_distrib_right algebra_simps pochhammer_Suc) finally show ?thesis . qed (insert assms, simp_all) lemma has_field_derivative_stirling_sum'_real [derivative_intros]: assumes "j > 0" "x > (0::real)" shows "(stirling_sum' j m has_field_derivative stirling_sum' (Suc j) m x) (at x)" proof (cases j) case (Suc j') from assms have [simp]: "x \ 0" by auto define c where "c = (\n. (-1) ^ Suc j * (bernoulli' n) * pochhammer (of_nat (Suc n)) j')" define T where "T = (\n x. c n * inverse x ^ (j + n))" define T' where "T' = (\n x. - (of_nat (j + n)) * c n * inverse x ^ (Suc (j + n)))" have "((\x. \k\m. T k x) has_field_derivative (\k\m. T' k x)) (at x)" using assms Suc by (intro DERIV_sum) (auto simp: T_def T'_def intro!: derivative_eq_intros simp: field_simps power_add [symmetric] simp del: of_nat_Suc power_Suc of_nat_add) also have "(\x. (\k\m. T k x)) = stirling_sum' j m" by (simp add: Suc T_def c_def stirling_sum'_def fun_eq_iff add_ac mult.assoc sum_distrib_left) also have "(\k\m. T' k x) = stirling_sum' (Suc j) m x" by (simp add: T'_def c_def Suc stirling_sum'_def sum_distrib_left sum_distrib_right algebra_simps pochhammer_Suc) finally show ?thesis . qed (insert assms, simp_all) lemma higher_deriv_stirling_sum_complex: "Re x > 0 \ (deriv ^^ i) (stirling_sum j m) x = stirling_sum (i + j) m x" proof (induction i arbitrary: x) case (Suc i) have "deriv ((deriv ^^ i) (stirling_sum j m)) x = deriv (stirling_sum (i + j) m) x" using eventually_nhds_in_open[of "{x. Re x > 0}" x] Suc.prems by (intro deriv_cong_ev refl) (auto elim!: eventually_mono simp: open_halfspace_Re_gt Suc.IH) also from Suc.prems have "\ = stirling_sum (Suc (i + j)) m x" by (intro DERIV_imp_deriv has_field_derivative_stirling_sum_complex) finally show ?case by simp qed simp_all definition Polygamma_approx :: "nat \ nat \ 'a \ 'a :: {real_normed_field, ln}" where "Polygamma_approx j m = (deriv ^^ j) (\x. (x - 1 / 2) * ln x - x + of_real (ln (2 * pi)) / 2 + stirling_sum 0 m x)" lemma Polygamma_approx_Suc: "Polygamma_approx (Suc j) m = deriv (Polygamma_approx j m)" by (simp add: Polygamma_approx_def) lemma Polygamma_approx_0: "Polygamma_approx 0 m x = (x - 1/2) * ln x - x + of_real (ln (2*pi)) / 2 + stirling_sum 0 m x" by (simp add: Polygamma_approx_def) lemma Polygamma_approx_1_complex: "Re x > 0 \ Polygamma_approx (Suc 0) m x = ln x - 1 / (2*x) + stirling_sum (Suc 0) m x" unfolding Polygamma_approx_Suc Polygamma_approx_0 by (intro DERIV_imp_deriv) (auto intro!: derivative_eq_intros elim!: nonpos_Reals_cases simp: field_simps) lemma Polygamma_approx_1_real: "x > (0 :: real) \ Polygamma_approx (Suc 0) m x = ln x - 1 / (2*x) + stirling_sum (Suc 0) m x" unfolding Polygamma_approx_Suc Polygamma_approx_0 by (intro DERIV_imp_deriv) (auto intro!: derivative_eq_intros elim!: nonpos_Reals_cases simp: field_simps) lemma stirling_sum_2_conv_stirling_sum'_1: fixes x :: "'a :: {real_div_algebra, field_char_0}" assumes "m > 0" "x \ 0" shows "stirling_sum' 1 m x = 1 / x + 1 / (2 * x^2) + stirling_sum 2 m x" proof - have pochhammer_2: "pochhammer (of_nat k) 2 = of_nat k * of_nat (Suc k)" for k by (simp add: pochhammer_Suc eval_nat_numeral add_ac) have "stirling_sum 2 m x = (\k = Suc 0.. = (\k=0.. = (\k=0.. = (\k\m. of_real (bernoulli' k) * inverse x ^ Suc k)" by (intro sum.cong) auto also have "\ = stirling_sum' 1 m x" by (simp add: stirling_sum'_def) finally show ?thesis by (simp add: add_ac) qed lemma Polygamma_approx_2_real: assumes "x > (0::real)" "m > 0" shows "Polygamma_approx (Suc (Suc 0)) m x = stirling_sum' 1 m x" proof - have "Polygamma_approx (Suc (Suc 0)) m x = deriv (Polygamma_approx (Suc 0) m) x" by (simp add: Polygamma_approx_Suc) also have "\ = deriv (\x. ln x - 1 / (2*x) + stirling_sum (Suc 0) m x) x" using eventually_nhds_in_open[of "{0<..}" x] assms by (intro deriv_cong_ev) (auto elim!: eventually_mono simp: Polygamma_approx_1_real) also have "\ = 1 / x + 1 / (2*x^2) + stirling_sum (Suc (Suc 0)) m x" using assms by (intro DERIV_imp_deriv) (auto intro!: derivative_eq_intros elim!: nonpos_Reals_cases simp: field_simps power2_eq_square) also have "\ = stirling_sum' 1 m x" using stirling_sum_2_conv_stirling_sum'_1[of m x] assms by (simp add: eval_nat_numeral) finally show ?thesis . qed lemma Polygamma_approx_2_complex: assumes "Re x > 0" "m > 0" shows "Polygamma_approx (Suc (Suc 0)) m x = stirling_sum' 1 m x" proof - have "Polygamma_approx (Suc (Suc 0)) m x = deriv (Polygamma_approx (Suc 0) m) x" by (simp add: Polygamma_approx_Suc) also have "\ = deriv (\x. ln x - 1 / (2*x) + stirling_sum (Suc 0) m x) x" using eventually_nhds_in_open[of "{s. Re s > 0}" x] assms by (intro deriv_cong_ev) (auto simp: open_halfspace_Re_gt elim!: eventually_mono simp: Polygamma_approx_1_complex) also have "\ = 1 / x + 1 / (2*x^2) + stirling_sum (Suc (Suc 0)) m x" using assms by (intro DERIV_imp_deriv) (auto intro!: derivative_eq_intros elim!: nonpos_Reals_cases simp: field_simps power2_eq_square) also have "\ = stirling_sum' 1 m x" using stirling_sum_2_conv_stirling_sum'_1[of m x] assms by (subst stirling_sum_2_conv_stirling_sum'_1) (auto simp: eval_nat_numeral) finally show ?thesis . qed lemma Polygamma_approx_ge_2_real: assumes "x > (0::real)" "m > 0" shows "Polygamma_approx (Suc (Suc j)) m x = stirling_sum' (Suc j) m x" using assms(1) proof (induction j arbitrary: x) case (0 x) with assms show ?case by (simp add: Polygamma_approx_2_real) next case (Suc j x) have "Polygamma_approx (Suc (Suc (Suc j))) m x = deriv (Polygamma_approx (Suc (Suc j)) m) x" by (simp add: Polygamma_approx_Suc) also have "\ = deriv (stirling_sum' (Suc j) m) x" using eventually_nhds_in_open[of "{0<..}" x] Suc.prems by (intro deriv_cong_ev refl) (auto elim!: eventually_mono simp: Suc.IH) also have "\ = stirling_sum' (Suc (Suc j)) m x" using Suc.prems by (intro DERIV_imp_deriv derivative_intros) simp_all finally show ?case . qed lemma Polygamma_approx_ge_2_complex: assumes "Re x > 0" "m > 0" shows "Polygamma_approx (Suc (Suc j)) m x = stirling_sum' (Suc j) m x" using assms(1) proof (induction j arbitrary: x) case (0 x) with assms show ?case by (simp add: Polygamma_approx_2_complex) next case (Suc j x) have "Polygamma_approx (Suc (Suc (Suc j))) m x = deriv (Polygamma_approx (Suc (Suc j)) m) x" by (simp add: Polygamma_approx_Suc) also have "\ = deriv (stirling_sum' (Suc j) m) x" using eventually_nhds_in_open[of "{x. Re x > 0}" x] Suc.prems by (intro deriv_cong_ev refl) (auto elim!: eventually_mono simp: Suc.IH open_halfspace_Re_gt) also have "\ = stirling_sum' (Suc (Suc j)) m x" using Suc.prems by (intro DERIV_imp_deriv derivative_intros) simp_all finally show ?case . qed lemma Polygamma_approx_complex_of_real: assumes "x > 0" "m > 0" shows "Polygamma_approx j m (complex_of_real x) = of_real (Polygamma_approx j m x)" proof (cases j) case 0 with assms show ?thesis by (simp add: Polygamma_approx_0 Ln_of_real stirling_sum_complex_of_real) next case [simp]: (Suc j') thus ?thesis proof (cases j') case 0 with assms show ?thesis by (simp add: Polygamma_approx_1_complex Polygamma_approx_1_real stirling_sum_complex_of_real Ln_of_real) next case (Suc j'') with assms show ?thesis by (simp add: Polygamma_approx_ge_2_complex Polygamma_approx_ge_2_real stirling_sum'_complex_of_real) qed qed lemma higher_deriv_Polygamma_approx [simp]: "(deriv ^^ j) (Polygamma_approx i m) = Polygamma_approx (j + i) m" by (simp add: Polygamma_approx_def funpow_add) lemma stirling_sum_holomorphic [holomorphic_intros]: "0 \ A \ stirling_sum j m holomorphic_on A" unfolding stirling_sum_def by (intro holomorphic_intros) auto lemma Polygamma_approx_holomorphic [holomorphic_intros]: "Polygamma_approx j m holomorphic_on {s. Re s > 0}" unfolding Polygamma_approx_def by (intro holomorphic_intros) (auto simp: open_halfspace_Re_gt elim!: nonpos_Reals_cases) lemma higher_deriv_lnGamma_stirling: assumes m: "m > 0" shows "(\x::real. (deriv ^^ j) ln_Gamma x - Polygamma_approx j m x) \ O(\x. 1 / x ^ (m + j))" proof - have "eventually (\x. \(deriv ^^ j) ln_Gamma x - Polygamma_approx j m x\ = inverse (real m) * \(deriv ^^ j) (stirling_integral m) x\) at_top" using eventually_gt_at_top[of "0::real"] proof eventually_elim case (elim x) note x = this have "\\<^sub>F y in nhds (complex_of_real x). y \ - \\<^sub>\\<^sub>0" using elim by (intro eventually_nhds_in_open) auto hence "(deriv ^^ j) (\x. ln_Gamma x - Polygamma_approx 0 m x) (complex_of_real x) = (deriv ^^ j) (\x. (-inverse (of_nat m)) * stirling_integral m x) (complex_of_real x)" using x m by (intro higher_deriv_cong_ev refl) (auto elim!: eventually_mono simp: ln_Gamma_stirling_complex Polygamma_approx_def field_simps open_halfspace_Re_gt stirling_sum_def) also have "\ = - inverse (of_nat m) * (deriv ^^ j) (stirling_integral m) (of_real x)" using x m by (intro higher_deriv_cmult[of _ "-\\<^sub>\\<^sub>0"] stirling_integral_holomorphic) (auto simp: open_halfspace_Re_gt) also have "(deriv ^^ j) (\x. ln_Gamma x - Polygamma_approx 0 m x) (complex_of_real x) = (deriv ^^ j) ln_Gamma (of_real x) - (deriv ^^ j) (Polygamma_approx 0 m) (of_real x)" using x by (intro higher_deriv_diff[of _ "{s. Re s > 0}"]) (auto intro!: holomorphic_intros elim!: nonpos_Reals_cases simp: open_halfspace_Re_gt) also have "(deriv ^^ j) (Polygamma_approx 0 m) (complex_of_real x) = of_real (Polygamma_approx j m x)" using x m by (simp add: Polygamma_approx_complex_of_real) also have "norm (- inverse (of_nat m) * (deriv ^^ j) (stirling_integral m) (complex_of_real x)) = inverse (real m) * \(deriv ^^ j) (stirling_integral m) x\" using x m by (simp add: norm_mult norm_inverse deriv_stirling_integral_complex_of_real) also have "(deriv ^^ j) ln_Gamma (complex_of_real x) = of_real ((deriv ^^ j) ln_Gamma x)" using x by (simp add: higher_deriv_ln_Gamma_complex_of_real) also have "norm (\ - of_real (Polygamma_approx j m x)) = \(deriv ^^ j) ln_Gamma x - Polygamma_approx j m x\" by (simp only: of_real_diff [symmetric] norm_of_real) finally show ?case . qed from bigthetaI_cong[OF this] m have "(\x::real. (deriv ^^ j) ln_Gamma x - Polygamma_approx j m x) \ \(\x. (deriv ^^ j) (stirling_integral m) x)" by simp also have "(\x::real. (deriv ^^ j) (stirling_integral m) x) \ O(\x. 1 / x ^ (m + j))" using m by (rule deriv_stirling_integral_real_bound) finally show ?thesis . qed lemma Polygamma_approx_1_real': assumes x: "(x::real) > 0" and m: "m > 0" shows "Polygamma_approx 1 m x = ln x - (\k = Suc 0..m. bernoulli' k * inverse x ^ k / real k)" proof - have "Polygamma_approx 1 m x = ln x - (1 / (2 * x) + (\k=Suc 0..k=Suc 0.. = (\k=0.. = (\k = Suc 0..m. bernoulli' k * inverse x ^ k / real k)" using assms by (subst sum.shift_bounds_Suc_ivl [symmetric]) (simp add: atLeastLessThanSuc_atLeastAtMost) finally show ?thesis . qed theorem assumes m: "m > 0" shows ln_Gamma_real_asymptotics: "(\x. ln_Gamma x - ((x - 1 / 2) * ln x - x + ln (2 * pi) / 2 + (\k = 1.. O(\x. 1 / x ^ m)" (is ?th1) and Digamma_real_asymptotics: "(\x. Digamma x - (ln x - (\k=1..m. bernoulli' k / real k / x ^ k))) \ O(\x. 1 / (x ^ Suc m))" (is ?th2) and Polygamma_real_asymptotics: "j > 0 \ (\x. Polygamma j x - (- 1) ^ Suc j * (\k\m. bernoulli' k * pochhammer (real (Suc k)) (j - 1) / x ^ (k + j))) \ O(\x. 1 / x ^ (m+j+1))" (is "_ \ ?th3") proof - define G :: "nat \ real \ real" where "G = (\m. if m = 0 then ln_Gamma else Polygamma (m - 1))" have *: "(\x. G j x - h x) \ O(\x. 1 / x ^ (m + j))" if "\x::real. x > 0 \ Polygamma_approx j m x = h x" for j h proof - have "(\x. G j x - h x) \ \(\x. (deriv ^^ j) ln_Gamma x - Polygamma_approx j m x)" (is "_ \ \(?f)") using that by (intro bigthetaI_cong) (auto intro: eventually_mono[OF eventually_gt_at_top[of "0::real"]] simp del: funpow.simps simp: higher_deriv_ln_Gamma_real G_def) also have "?f \ O(\x::real. 1 / x ^ (m + j))" using m by (rule higher_deriv_lnGamma_stirling) finally show ?thesis . qed note [[simproc del: simplify_landau_sum]] from *[OF Polygamma_approx_0] assms show ?th1 by (simp add: G_def Polygamma_approx_0 stirling_sum_def field_simps) from *[OF Polygamma_approx_1_real'] assms show ?th2 by (simp add: G_def field_simps) assume j: "j > 0" from *[OF Polygamma_approx_ge_2_real, of "j - 1"] assms j show ?th3 by (simp add: G_def stirling_sum'_def power_add power_diff field_simps) qed subsection \Asymptotics of the complex Gamma function\ text \ The \m\-th order remainder of Stirling's formula for $\log\Gamma$ is $O(s^{-m})$ uniformly over - any complex cone $\text{arg}(z) \leq \alpha$, $z\neq 0$ for any angle + any complex cone $\text{Arg}(z) \leq \alpha$, $z\neq 0$ for any angle $\alpha\in(0, \pi)$. This means that there is bounded by $c z^{-m}$ for some constant $c$ for all $z$ in this cone. \ context fixes F and \ assumes \: "\ \ {0<.. principal (complex_cone' \ - {0})" begin lemma stirling_integral_bigo: fixes m :: nat assumes m: "m > 0" shows "stirling_integral m \ O[F](\s. 1 / s ^ m)" proof - obtain c where c: "\s. s \ complex_cone' \ - {0} \ norm (stirling_integral m s) \ c / norm s ^ m" using stirling_integral_bound'[OF \m > 0\ \] by blast have "0 \ norm (stirling_integral m 1 :: complex)" by simp also have "\ \ c" using c[of 1] \ by simp finally have "c \ 0" . have "eventually (\s. s \ complex_cone' \ - {0}) F" unfolding F_def by (auto simp: eventually_principal) hence "eventually (\s. norm (stirling_integral m s) \ c * norm (1 / s ^ m)) F" by eventually_elim (use c in \simp add: norm_divide norm_power\) thus "stirling_integral m \ O[F](\s. 1 / s ^ m)" by (intro bigoI[of _ c]) auto qed end text \ The following is a more explicit statement of this: \ theorem ln_Gamma_complex_asymptotics_explicit: fixes m :: nat and \ :: real assumes "m > 0" and "\ \ {0<.. complex" where "\s::complex. s \ \\<^sub>\\<^sub>0 \ ln_Gamma s = (s - 1/2) * ln s - s + ln (2 * pi) / 2 + (\k=1..s. s \ 0 \ \arg s\ \ \ \ norm (R s) \ C / norm s ^ m" + and "\s. s \ 0 \ \Arg s\ \ \ \ norm (R s) \ C / norm s ^ m" proof - obtain c where c: "\s. s \ complex_cone' \ - {0} \ norm (stirling_integral m s) \ c / norm s ^ m" using stirling_integral_bound'[OF assms] by blast have "0 \ norm (stirling_integral m 1 :: complex)" by simp also have "\ \ c" using c[of 1] assms by simp finally have "c \ 0" . define R where "R = (\s::complex. stirling_integral m s / of_nat m)" show ?thesis proof (rule that) from ln_Gamma_stirling_complex[of _ m] assms show "\s::complex. s \ \\<^sub>\\<^sub>0 \ ln_Gamma s = (s - 1 / 2) * ln s - s + ln (2 * pi) / 2 + (\k=1..s. s \ 0 \ \arg s\ \ \ \ cmod (R s) \ c / real m / cmod s ^ m" + show "\s. s \ 0 \ \Arg s\ \ \ \ cmod (R s) \ c / real m / cmod s ^ m" proof (safe, goal_cases) case (1 s) show ?case using 1 c[of s] assms by (auto simp: complex_cone_altdef abs_le_iff R_def norm_divide field_simps) qed qed qed text \ Lastly, we can also derive the asymptotics of $\Gamma$ itself: \[\Gamma(z) \sim \sqrt{2\pi / z} \left(\frac{z}{e}\right)^z\] - uniformly for $|z|\to\infty$ within the cone $\text{arg}(z) \leq \alpha$ for $\alpha\in(0,\pi)$: + uniformly for $|z|\to\infty$ within the cone $\text{Arg}(z) \leq \alpha$ for $\alpha\in(0,\pi)$: \ context fixes F and \ assumes \: "\ \ {0<.. inf at_infinity (principal (complex_cone' \))" begin lemma Gamma_complex_asymp_equiv: "Gamma \[F] (\s. sqrt (2 * pi) * (s / exp 1) powr s / s powr (1 / 2))" proof - define I :: "complex \ complex" where "I = stirling_integral 1" have "eventually (\s. s \ complex_cone' \) F" by (auto simp: eventually_inf_principal F_def) moreover have "eventually (\s. s \ 0) F" unfolding F_def eventually_inf_principal using eventually_not_equal_at_infinity by eventually_elim auto ultimately have "eventually (\s. Gamma s = sqrt (2 * pi) * (s / exp 1) powr s / s powr (1 / 2) / exp (I s)) F" proof eventually_elim case (elim s) from elim have s': "s \ \\<^sub>\\<^sub>0" using complex_cone_inter_nonpos_Reals[of "-\" \] \ by auto from elim have [simp]: "s \ 0" by auto from s' have "Gamma s = exp (ln_Gamma s)" unfolding Gamma_complex_altdef using nonpos_Ints_subset_nonpos_Reals by auto also from s' have "ln_Gamma s = (s-1/2) * Ln s - s + complex_of_real (ln (2 * pi) / 2) - I s" by (subst ln_Gamma_stirling_complex[of _ 1]) (simp_all add: exp_add exp_diff I_def) also have "exp \ = exp ((s - 1 / 2) * Ln s) / exp s * exp (complex_of_real (ln (2 * pi) / 2)) / exp (I s)" unfolding exp_diff exp_add by (simp add: exp_diff exp_add) also have "exp ((s - 1 / 2) * Ln s) = s powr (s - 1 / 2)" by (simp add: powr_def) also have "exp (complex_of_real (ln (2 * pi) / 2)) = sqrt (2 * pi)" by (subst exp_of_real) (auto simp: powr_def simp flip: powr_half_sqrt) also have "exp s = exp 1 powr s" by (simp add: powr_def) also have "s powr (s - 1 / 2) / exp 1 powr s = (s powr s / exp 1 powr s) / s powr (1/2)" by (subst powr_diff) auto also have *: "Ln (s / exp 1) = Ln s - 1" using Ln_divide_of_real[of "exp 1" s] by (simp flip: exp_of_real) hence "s powr s / exp 1 powr s = (s / exp 1) powr s" unfolding powr_def by (subst *) (auto simp: exp_diff field_simps) finally show "Gamma s = sqrt (2 * pi) * (s / exp 1) powr s / s powr (1 / 2) / exp (I s)" by (simp add: algebra_simps) qed hence "Gamma \[F] (\s. sqrt (2 * pi) * (s / exp 1) powr s / s powr (1 / 2) / exp (I s))" by (rule asymp_equiv_refl_ev) also have "\ \[F] (\s. sqrt (2 * pi) * (s / exp 1) powr s / s powr (1 / 2) / 1)" proof (intro asymp_equiv_intros) have "F \ principal (complex_cone' \ - {0})" unfolding le_principal F_def eventually_inf_principal using eventually_not_equal_at_infinity by eventually_elim auto moreover have "I \ O[principal (complex_cone' \ - {0})](\s. 1 / s)" using stirling_integral_bigo[of \ 1] \ unfolding F_def by (simp add: I_def) ultimately have "I \ O[F](\s. 1 / s)" by (rule landau_o.big.filter_mono) also have "(\s. 1 / s) \ o[F](\s. 1)" proof (rule landau_o.smallI) fix c :: real assume c: "c > 0" hence "eventually (\z::complex. norm z \ 1 / c) at_infinity" by (auto simp: eventually_at_infinity) moreover have "eventually (\z::complex. z \ 0) at_infinity" by (rule eventually_not_equal_at_infinity) ultimately show "eventually (\z::complex. norm (1 / z) \ c * norm (1 :: complex)) F" unfolding F_def eventually_inf_principal by eventually_elim (use \c > 0\ in \auto simp: norm_divide field_simps\) qed finally have "I \ o[F](\s. 1)" . from smalloD_tendsto[OF this] have [tendsto_intros]: "(I \ 0) F" by simp show "(\x. exp (I x)) \[F] (\x. 1)" by (rule asymp_equivI' tendsto_eq_intros refl | simp)+ qed finally show ?thesis by simp qed end end