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") + (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") + 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'" + 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'" + 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) 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" + 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" + 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 + 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_Lines.thy b/thys/Poincare_Disc/Poincare_Lines.thy --- a/thys/Poincare_Disc/Poincare_Lines.thy +++ b/thys/Poincare_Disc/Poincare_Lines.thy @@ -1,1878 +1,1878 @@ (* ------------------------------------------------------------------ *) section \H-lines in the Poincar\'e model\ (* ------------------------------------------------------------------ *) theory Poincare_Lines imports Complex_Geometry.Unit_Circle_Preserving_Moebius Complex_Geometry.Circlines_Angle begin (* ------------------------------------------------------------------ *) subsection \Definition and basic properties of h-lines\ (* ------------------------------------------------------------------ *) text \H-lines in the Poincar\'e model are either line segments passing trough the origin or segments (within the unit disc) of circles that are perpendicular to the unit circle. Algebraically these are circlines that are represented by Hermitean matrices of the form $$H = \left( \begin{array}{cc} A & B\\ \overline{B} & A \end{array} \right),$$ for $A \in \mathbb{R}$, and $B \in \mathbb{C}$, and $|B|^2 > A^2$, where the circline equation is the usual one: $z^*Hz = 0$, for homogenous coordinates $z$.\ definition is_poincare_line_cmat :: "complex_mat \ bool" where [simp]: "is_poincare_line_cmat H \ (let (A, B, C, D) = H in hermitean (A, B, C, D) \ A = D \ (cmod B)\<^sup>2 > (cmod A)\<^sup>2)" lift_definition is_poincare_line_clmat :: "circline_mat \ bool" is is_poincare_line_cmat done text \We introduce the predicate that checks if a given complex matrix is a matrix of a h-line in the Poincar\'e model, and then by means of the lifting package lift it to the type of non-zero Hermitean matrices, and then to circlines (that are equivalence classes of such matrices).\ lift_definition is_poincare_line :: "circline \ bool" is is_poincare_line_clmat proof (transfer, transfer) fix H1 H2 :: complex_mat assume hh: "hermitean H1 \ H1 \ mat_zero" "hermitean H2 \ H2 \ mat_zero" assume "circline_eq_cmat H1 H2" thus "is_poincare_line_cmat H1 \ is_poincare_line_cmat H2" using hh by (cases H1, cases H2) (auto simp: norm_mult power_mult_distrib) qed lemma is_poincare_line_mk_circline: assumes "(A, B, C, D) \ hermitean_nonzero" shows "is_poincare_line (mk_circline A B C D) \ (cmod B)\<^sup>2 > (cmod A)\<^sup>2 \ A = D" using assms by (transfer, transfer, auto simp add: Let_def) text\Abstract characterisation of @{term is_poincare_line} predicate: H-lines in the Poincar\'e model are real circlines (circlines with the negative determinant) perpendicular to the unit circle.\ lemma is_poincare_line_iff: shows "is_poincare_line H \ circline_type H = -1 \ perpendicular H unit_circle" unfolding perpendicular_def proof (simp, transfer, transfer) fix H assume hh: "hermitean H \ H \ mat_zero" obtain A B C D where *: "H = (A, B, C, D)" by (cases H, auto) have **: "is_real A" "is_real D" "C = cnj B" using hh * hermitean_elems by auto hence "(Re A = Re D \ cmod A * cmod A < cmod B * cmod B) = (Re A * Re D < Re B * Re B + Im B * Im B \ (Re D = Re A \ Re A * Re D = Re B * Re B + Im B * Im B))" using * by (smt cmod_power2 power2_eq_square zero_power2)+ thus "is_poincare_line_cmat H \ circline_type_cmat H = - 1 \ cos_angle_cmat (of_circline_cmat H) unit_circle_cmat = 0" using * ** by (auto simp add: sgn_1_neg complex_eq_if_Re_eq cmod_square power2_eq_square simp del: pos_oriented_cmat_def) qed text\The @{term x_axis} is an h-line.\ lemma is_poincare_line_x_axis [simp]: shows "is_poincare_line x_axis" by (transfer, transfer) (auto simp add: hermitean_def mat_adj_def mat_cnj_def) text\The @{term unit_circle} is not an h-line.\ lemma not_is_poincare_line_unit_circle [simp]: shows "\ is_poincare_line unit_circle" by (transfer, transfer, simp) (* ------------------------------------------------------------------ *) subsubsection \Collinear points\ (* ------------------------------------------------------------------ *) text\Points are collinear if they all belong to an h-line. \ definition poincare_collinear :: "complex_homo set \ bool" where "poincare_collinear S \ (\ p. is_poincare_line p \ S \ circline_set p)" (* ------------------------------------------------------------------ *) subsubsection \H-lines and inversion\ (* ------------------------------------------------------------------ *) text\Every h-line in the Poincar\'e model contains the inverse (wrt.~the unit circle) of each of its points (note that at most one of them belongs to the unit disc).\ lemma is_poincare_line_inverse_point: assumes "is_poincare_line H" "u \ circline_set H" shows "inversion u \ circline_set H" using assms unfolding is_poincare_line_iff circline_set_def perpendicular_def inversion_def apply simp proof (transfer, transfer) fix u H assume hh: "hermitean H \ H \ mat_zero" "u \ vec_zero" and aa: "circline_type_cmat H = - 1 \ cos_angle_cmat (of_circline_cmat H) unit_circle_cmat = 0" "on_circline_cmat_cvec H u" obtain A B C D u1 u2 where *: "H = (A, B, C, D)" "u = (u1, u2)" by (cases H, cases u, auto) have "is_real A" "is_real D" "C = cnj B" using * hh hermitean_elems by auto moreover have "A = D" using aa(1) * \is_real A\ \is_real D\ by (auto simp del: pos_oriented_cmat_def simp add: complex.expand split: if_split_asm) thus "on_circline_cmat_cvec H (conjugate_cvec (reciprocal_cvec u))" using aa(2) * by (simp add: vec_cnj_def field_simps) qed text\Every h-line in the Poincar\'e model and is invariant under unit circle inversion.\ lemma circline_inversion_poincare_line: assumes "is_poincare_line H" shows "circline_inversion H = H" proof- obtain u v w where *: "u \ v" "v \ w" "u \ w" "{u, v, w} \ circline_set H" using assms is_poincare_line_iff[of H] using circline_type_neg_card_gt3[of H] by auto hence "{inversion u, inversion v, inversion w} \ circline_set (circline_inversion H)" "{inversion u, inversion v, inversion w} \ circline_set H" using is_poincare_line_inverse_point[OF assms] by auto thus ?thesis using * unique_circline_set[of "inversion u" "inversion v" "inversion w"] by (metis insert_subset inversion_involution) qed (* ------------------------------------------------------------------ *) subsubsection \Classification of h-lines into Euclidean segments and circles\ (* ------------------------------------------------------------------ *) text\If an h-line contains zero, than it also contains infinity (the inverse point of zero) and is by definition an Euclidean line.\ lemma is_poincare_line_trough_zero_trough_infty [simp]: assumes "is_poincare_line l" and "0\<^sub>h \ circline_set l" shows "\\<^sub>h \ circline_set l" using is_poincare_line_inverse_point[OF assms] by simp lemma is_poincare_line_trough_zero_is_line: assumes "is_poincare_line l" and "0\<^sub>h \ circline_set l" shows "is_line l" using assms using inf_in_circline_set is_poincare_line_trough_zero_trough_infty by blast text\If an h-line does not contain zero, than it also does not contain infinity (the inverse point of zero) and is by definition an Euclidean circle.\ lemma is_poincare_line_not_trough_zero_not_trough_infty [simp]: assumes "is_poincare_line l" assumes "0\<^sub>h \ circline_set l" shows "\\<^sub>h \ circline_set l" using assms using is_poincare_line_inverse_point[OF assms(1), of "\\<^sub>h"] by auto lemma is_poincare_line_not_trough_zero_is_circle: assumes "is_poincare_line l" "0\<^sub>h \ circline_set l" shows "is_circle l" using assms using inf_in_circline_set is_poincare_line_not_trough_zero_not_trough_infty by auto (* ------------------------------------------------------------------ *) subsubsection\Points on h-line\ (* ------------------------------------------------------------------ *) text\Each h-line in the Poincar\'e model contains at least two different points within the unit disc.\ text\First we prove an auxiliary lemma.\ lemma ex_is_poincare_line_points': assumes i12: "i1 \ circline_set H \ unit_circle_set" "i2 \ circline_set H \ unit_circle_set" "i1 \ i2" assumes a: "a \ circline_set H" "a \ unit_circle_set" shows "\ b. b \ i1 \ b \ i2 \ b \ a \ b \ inversion a \ b \ circline_set H" proof- have "inversion a \ unit_circle_set" using \a \ unit_circle_set\ unfolding unit_circle_set_def circline_set_def by (metis inversion_id_iff_on_unit_circle inversion_involution mem_Collect_eq) have "a \ inversion a" using \a \ unit_circle_set\ inversion_id_iff_on_unit_circle[of a] unfolding unit_circle_set_def circline_set_def by auto have "a \ i1" "a \ i2" "inversion a \ i1" "inversion a \ i2" using assms \inversion a \ unit_circle_set\ by auto then obtain b where cr2: "cross_ratio b i1 a i2 = of_complex 2" using \i1 \ i2\ using ex_cross_ratio[of i1 a i2] by blast have distinct_b: "b \ i1" "b \ i2" "b \ a" using \i1 \ i2\ \a \ i1\ \a \ i2\ using ex1_cross_ratio[of i1 a i2] using cross_ratio_0[of i1 a i2] cross_ratio_1[of i1 a i2] cross_ratio_inf[of i1 i2 a] using cr2 by auto hence "b \ circline_set H" using assms four_points_on_circline_iff_cross_ratio_real[of b i1 a i2] cr2 using unique_circline_set[of i1 i2 a] by auto moreover have "b \ inversion a" proof (rule ccontr) assume *: "\ ?thesis" have "inversion i1 = i1" "inversion i2 = i2" using i12 unfolding unit_circle_set_def by auto hence "cross_ratio (inversion a) i1 a i2 = cross_ratio a i1 (inversion a) i2" using * cross_ratio_inversion[of i1 a i2 b] \a \ i1\ \a \ i2\ \i1 \ i2\ \b \ i1\ using four_points_on_circline_iff_cross_ratio_real[of b i1 a i2] using i12 distinct_b conjugate_id_iff[of "cross_ratio b i1 a i2"] using i12 a \b \ circline_set H\ by auto hence "cross_ratio (inversion a) i1 a i2 \ of_complex 2" using cross_ratio_commute_13[of "inversion a" i1 a i2] using reciprocal_id_iff using of_complex_inj by force thus False using * cr2 by simp qed ultimately show ?thesis using assms \b \ i1\ \b \ i2\ \b \ a\ by auto qed text\Now we can prove the statement.\ lemma ex_is_poincare_line_points: assumes "is_poincare_line H" shows "\ u v. u \ unit_disc \ v \ unit_disc \ u \ v \ {u, v} \ circline_set H" proof- obtain u v w where *: "u \ v" "v \ w" "u \ w" "{u, v, w} \ circline_set H" using assms is_poincare_line_iff[of H] using circline_type_neg_card_gt3[of H] by auto have "\ {u, v, w} \ unit_circle_set" using unique_circline_set[of u v w] * by (metis assms insert_subset not_is_poincare_line_unit_circle unit_circle_set_def) hence "H \ unit_circle" unfolding unit_circle_set_def using * by auto show ?thesis proof (cases "(u \ unit_disc \ v \ unit_disc) \ (u \ unit_disc \ w \ unit_disc) \ (v \ unit_disc \ w \ unit_disc)") case True thus ?thesis using * by auto next case False have "\ a b. a \ b \ a \ inversion b \ a \ circline_set H \ b \ circline_set H \ a \ unit_circle_set \ b \ unit_circle_set" proof (cases "(u \ unit_circle_set \ v \ unit_circle_set) \ (u \ unit_circle_set \ w \ unit_circle_set) \ (v \ unit_circle_set \ w \ unit_circle_set)") case True then obtain i1 i2 a where *: "i1 \ unit_circle_set \ circline_set H" "i2 \ unit_circle_set \ circline_set H" "a \ circline_set H" "a \ unit_circle_set" "i1 \ i2" "i1 \ a" "i2 \ a" using * \\ {u, v, w} \ unit_circle_set\ by auto then obtain b where "b \ circline_set H" "b \ i1" "b \ i2" "b \ a" "b \ inversion a" using ex_is_poincare_line_points'[of i1 H i2 a] by blast hence "b \ unit_circle_set" using * \H \ unit_circle\ unique_circline_set[of i1 i2 b] unfolding unit_circle_set_def by auto thus ?thesis using * \b \ circline_set H\ \b \ a\ \b \ inversion a\ by auto next case False then obtain f g h where *: "f \ g" "f \ circline_set H" "f \ unit_circle_set" "g \ circline_set H" "g \ unit_circle_set" "h \ circline_set H" "h \ f" "h \ g" using * by auto show ?thesis proof (cases "f = inversion g") case False thus ?thesis using * by auto next case True show ?thesis proof (cases "h \ unit_circle_set") case False thus ?thesis using * \f = inversion g\ by auto next case True obtain m where cr2: "cross_ratio m h f g = of_complex 2" using ex_cross_ratio[of h f g] * \f \ g\ \h \ f\ \h \ g\ by auto hence "m \ h" "m \ f" "m \ g" using \h \ f\ \h \ g\ \f \ g\ using ex1_cross_ratio[of h f g] using cross_ratio_0[of h f g] cross_ratio_1[of h f g] cross_ratio_inf[of h g f] using cr2 by auto hence "m \ circline_set H" using four_points_on_circline_iff_cross_ratio_real[of m h f g] cr2 using \h \ f\ \h \ g\ \f \ g\ * using unique_circline_set[of h f g] by auto show ?thesis proof (cases "m \ unit_circle_set") case False thus ?thesis using \m \ f\ \m \ g\ \f = inversion g\ * \m \ circline_set H\ by auto next case True then obtain n where "n \ h" "n \ m" "n \ f" "n \ inversion f" "n \ circline_set H" using ex_is_poincare_line_points'[of h H m f] * \m \ circline_set H\ \h \ unit_circle_set\ \m \ h\ by auto hence "n \ unit_circle_set" using * \H \ unit_circle\ unique_circline_set[of m n h] using \m \ h\ \m \ unit_circle_set\ \h \ unit_circle_set\ \m \ circline_set H\ unfolding unit_circle_set_def by auto thus ?thesis using * \n \ circline_set H\ \n \ f\ \n \ inversion f\ by auto qed qed qed qed then obtain a b where ab: "a \ b" "a \ inversion b" "a \ circline_set H" "b \ circline_set H" "a \ unit_circle_set" "b \ unit_circle_set" by blast have "\ x. x \ circline_set H \ x \ unit_circle_set \ (\ x'. x' \ circline_set H \ unit_disc \ (x' = x \ x' = inversion x))" proof safe fix x assume x: "x \ circline_set H" "x \ unit_circle_set" show "\ x'. x' \ circline_set H \ unit_disc \ (x' = x \ x' = inversion x)" proof (cases "x \ unit_disc") case True thus ?thesis using x by auto next case False hence "x \ unit_disc_compl" using x in_on_out_univ[of "ounit_circle"] unfolding unit_circle_set_def unit_disc_def unit_disc_compl_def by auto hence "inversion x \ unit_disc" using inversion_unit_disc_compl by blast thus ?thesis using is_poincare_line_inverse_point[OF assms, of x] x by auto qed qed then obtain a' b' where *: "a' \ circline_set H" "a' \ unit_disc" "b' \ circline_set H" "b' \ unit_disc" and **: "a' = a \ a' = inversion a" "b' = b \ b' = inversion b" using ab by blast have "a' \ b'" using \a \ b\ \a \ inversion b\ ** * by (metis inversion_involution) thus ?thesis using * by auto qed qed (* ------------------------------------------------------------------ *) subsubsection \H-line uniqueness\ (* ------------------------------------------------------------------ *) text\There is no more than one h-line that contains two different h-points (in the disc).\ lemma unique_is_poincare_line: assumes in_disc: "u \ unit_disc" "v \ unit_disc" "u \ v" assumes pl: "is_poincare_line l1" "is_poincare_line l2" assumes on_l: "{u, v} \ circline_set l1 \ circline_set l2" shows "l1 = l2" proof- have "u \ inversion u" "v \ inversion u" using in_disc using inversion_noteq_unit_disc[of u v] using inversion_noteq_unit_disc[of u u] by auto thus ?thesis using on_l using unique_circline_set[of u "inversion u" "v"] \u \ v\ using is_poincare_line_inverse_point[of l1 u] using is_poincare_line_inverse_point[of l2 u] using pl by auto qed text\For the rest of our formalization it is often useful to consider points on h-lines that are not within the unit disc. Many lemmas in the rest of this section will have such generalizations.\ text\There is no more than one h-line that contains two different and not mutually inverse points (not necessary in the unit disc).\ lemma unique_is_poincare_line_general: assumes different: "u \ v" "u \ inversion v" assumes pl: "is_poincare_line l1" "is_poincare_line l2" assumes on_l: "{u, v} \ circline_set l1 \ circline_set l2" shows "l1 = l2" proof (cases "u \ inversion u") case True thus ?thesis using unique_circline_set[of u "inversion u" "v"] using assms using is_poincare_line_inverse_point by force next case False show ?thesis proof (cases "v \ inversion v") case True thus ?thesis using unique_circline_set[of u "inversion v" "v"] using assms using is_poincare_line_inverse_point by force next case False have "on_circline unit_circle u" "on_circline unit_circle v" using \\ u \ inversion u\ \\ v \ inversion v\ using inversion_id_iff_on_unit_circle by fastforce+ thus ?thesis using pl on_l \u \ v\ unfolding circline_set_def apply simp proof (transfer, transfer, safe) fix u1 u2 v1 v2 A1 B1 C1 D1 A2 B2 C2 D2 :: complex let ?u = "(u1, u2)" and ?v = "(v1, v2)" and ?H1 = "(A1, B1, C1, D1)" and ?H2 = "(A2, B2, C2, D2)" assume *: "?u \ vec_zero" "?v \ vec_zero" "on_circline_cmat_cvec unit_circle_cmat ?u" "on_circline_cmat_cvec unit_circle_cmat ?v" "is_poincare_line_cmat ?H1" "is_poincare_line_cmat ?H2" "hermitean ?H1" "?H1 \ mat_zero" "hermitean ?H2" "?H2 \ mat_zero" "on_circline_cmat_cvec ?H1 ?u" "on_circline_cmat_cvec ?H1 ?v" "on_circline_cmat_cvec ?H2 ?u" "on_circline_cmat_cvec ?H2 ?v" "\ (u1, u2) \\<^sub>v (v1, v2)" have **: "A1 = D1" "A2 = D2" "C1 = cnj B1" "C2 = cnj B2" "is_real A1" "is_real A2" using \is_poincare_line_cmat ?H1\ \is_poincare_line_cmat ?H2\ using \hermitean ?H1\ \?H1 \ mat_zero\ \hermitean ?H2\ \?H2 \ mat_zero\ using hermitean_elems by auto have uv: "u1 \ 0" "u2 \ 0" "v1 \ 0" "v2 \ 0" using *(1-4) by (auto simp add: vec_cnj_def) have u: "cor ((Re (u1/u2))\<^sup>2) + cor ((Im (u1/u2))\<^sup>2) = 1" using \on_circline_cmat_cvec unit_circle_cmat ?u\ uv apply (subst of_real_add[symmetric]) apply (subst complex_mult_cnj[symmetric]) apply (simp add: vec_cnj_def mult.commute) done have v: "cor ((Re (v1/v2))\<^sup>2) + cor ((Im (v1/v2))\<^sup>2) = 1" using \on_circline_cmat_cvec unit_circle_cmat ?v\ uv apply (subst of_real_add[symmetric]) apply (subst complex_mult_cnj[symmetric]) apply (simp add: vec_cnj_def mult.commute) done have "A1 * (cor ((Re (u1/u2))\<^sup>2) + cor ((Im (u1/u2))\<^sup>2) + 1) + cor (Re B1) * cor(2 * Re (u1/u2)) + cor (Im B1) * cor(2 * Im (u1/u2)) = 0" "A2 * (cor ((Re (u1/u2))\<^sup>2) + cor ((Im (u1/u2))\<^sup>2) + 1) + cor (Re B2) * cor(2 * Re (u1/u2)) + cor (Im B2) * cor(2 * Im (u1/u2)) = 0" "A1 * (cor ((Re (v1/v2))\<^sup>2) + cor ((Im (v1/v2))\<^sup>2) + 1) + cor (Re B1) * cor(2 * Re (v1/v2)) + cor (Im B1) * cor(2 * Im (v1/v2)) = 0" "A2 * (cor ((Re (v1/v2))\<^sup>2) + cor ((Im (v1/v2))\<^sup>2) + 1) + cor (Re B2) * cor(2 * Re (v1/v2)) + cor (Im B2) * cor(2 * Im (v1/v2)) = 0" using circline_equation_quadratic_equation[of A1 "u1/u2" B1 D1 "Re (u1/u2)" "Im (u1 / u2)" "Re B1" "Im B1"] using circline_equation_quadratic_equation[of A2 "u1/u2" B2 D2 "Re (u1/u2)" "Im (u1 / u2)" "Re B2" "Im B2"] using circline_equation_quadratic_equation[of A1 "v1/v2" B1 D1 "Re (v1/v2)" "Im (v1 / v2)" "Re B1" "Im B1"] using circline_equation_quadratic_equation[of A2 "v1/v2" B2 D2 "Re (v1/v2)" "Im (v1 / v2)" "Re B2" "Im B2"] using \on_circline_cmat_cvec ?H1 ?u\ \on_circline_cmat_cvec ?H2 ?u\ using \on_circline_cmat_cvec ?H1 ?v\ \on_circline_cmat_cvec ?H2 ?v\ using ** uv by (simp_all add: vec_cnj_def field_simps) hence "A1 + cor (Re B1) * cor(Re (u1/u2)) + cor (Im B1) * cor(Im (u1/u2)) = 0" "A1 + cor (Re B1) * cor(Re (v1/v2)) + cor (Im B1) * cor(Im (v1/v2)) = 0" "A2 + cor (Re B2) * cor(Re (u1/u2)) + cor (Im B2) * cor(Im (u1/u2)) = 0" "A2 + cor (Re B2) * cor(Re (v1/v2)) + cor (Im B2) * cor(Im (v1/v2)) = 0" using u v by simp_all algebra+ hence "cor (Re A1 + Re B1 * Re (u1/u2) + Im B1 * Im (u1/u2)) = 0" "cor (Re A2 + Re B2 * Re (u1/u2) + Im B2 * Im (u1/u2)) = 0" "cor (Re A1 + Re B1 * Re (v1/v2) + Im B1 * Im (v1/v2)) = 0" "cor (Re A2 + Re B2 * Re (v1/v2) + Im B2 * Im (v1/v2)) = 0" using \is_real A1\ \is_real A2\ by simp_all hence "Re A1 + Re B1 * Re (u1/u2) + Im B1 * Im (u1/u2) = 0" "Re A1 + Re B1 * Re (v1/v2) + Im B1 * Im (v1/v2) = 0" "Re A2 + Re B2 * Re (u1/u2) + Im B2 * Im (u1/u2) = 0" "Re A2 + Re B2 * Re (v1/v2) + Im B2 * Im (v1/v2) = 0" using of_real_eq_0_iff by blast+ moreover have "Re(u1/u2) \ Re(v1/v2) \ Im(u1/u2) \ Im(v1/v2)" proof (rule ccontr) assume "\ ?thesis" hence "u1/u2 = v1/v2" using complex_eqI by blast thus False using uv \\ (u1, u2) \\<^sub>v (v1, v2)\ using "*"(1) "*"(2) complex_cvec_eq_mix[OF *(1) *(2)] by (auto simp add: field_simps) qed moreover have "Re A1 \ 0 \ Re B1 \ 0 \ Im B1 \ 0" using \?H1 \ mat_zero\ ** by (metis complex_cnj_zero complex_of_real_Re mat_zero_def of_real_0) ultimately obtain k where k: "Re A2 = k * Re A1" "Re B2 = k * Re B1" "Im B2 = k * Im B1" using linear_system_homogenous_3_2[of "\x y z. 1 * x + Re (u1 / u2) * y + Im (u1 / u2) * z" 1 "Re (u1/u2)" "Im (u1/u2)" "\x y z. 1 * x + Re (v1 / v2) * y + Im (v1 / v2) * z" 1 "Re (v1/v2)" "Im (v1/v2)" "Re A2" "Re B2" "Im B2" "Re A1" "Re B1" "Im B1"] by (auto simp add: field_simps) have "Re A2 \ 0 \ Re B2 \ 0 \ Im B2 \ 0" using \?H2 \ mat_zero\ ** by (metis complex_cnj_zero complex_of_real_Re mat_zero_def of_real_0) hence "k \ 0" using k by auto show "circline_eq_cmat ?H1 ?H2" using ** k \k \ 0\ by (auto simp add: vec_cnj_def) (rule_tac x="k" in exI, auto simp add: complex.expand) qed qed qed text \The only h-line that goes trough zero and a non-zero point on the x-axis is the x-axis.\ lemma is_poincare_line_0_real_is_x_axis: assumes "is_poincare_line l" "0\<^sub>h \ circline_set l" "x \ circline_set l \ circline_set x_axis" "x \ 0\<^sub>h" "x \ \\<^sub>h" shows "l = x_axis" using assms using is_poincare_line_trough_zero_trough_infty[OF assms(1-2)] using unique_circline_set[of x "0\<^sub>h" "\\<^sub>h"] by auto text \The only h-line that goes trough zero and a non-zero point on the y-axis is the y-axis.\ lemma is_poincare_line_0_imag_is_y_axis: assumes "is_poincare_line l" "0\<^sub>h \ circline_set l" "y \ circline_set l \ circline_set y_axis" "y \ 0\<^sub>h" "y \ \\<^sub>h" shows "l = y_axis" using assms using is_poincare_line_trough_zero_trough_infty[OF assms(1-2)] using unique_circline_set[of y "0\<^sub>h" "\\<^sub>h"] by auto (* ------------------------------------------------------------------ *) subsubsection\H-isometries preserve h-lines\ (* ------------------------------------------------------------------ *) text\\emph{H-isometries} are defined as homographies (actions of Möbius transformations) and antihomographies (compositions of actions of Möbius transformations with conjugation) that fix the unit disc (map it onto itself). They also map h-lines onto h-lines\ text\We prove a bit more general lemma that states that all Möbius transformations that fix the unit circle (not necessary the unit disc) map h-lines onto h-lines\ lemma unit_circle_fix_preserve_is_poincare_line [simp]: assumes "unit_circle_fix M" "is_poincare_line H" shows "is_poincare_line (moebius_circline M H)" using assms unfolding is_poincare_line_iff proof (safe) let ?H' = "moebius_ocircline M (of_circline H)" let ?U' = "moebius_ocircline M ounit_circle" assume ++: "unit_circle_fix M" "perpendicular H unit_circle" have ounit: "ounit_circle = moebius_ocircline M ounit_circle \ ounit_circle = moebius_ocircline M (opposite_ocircline ounit_circle)" using ++(1) unit_circle_fix_iff[of M] by (simp add: inj_of_ocircline moebius_circline_ocircline) show "perpendicular (moebius_circline M H) unit_circle" proof (cases "pos_oriented ?H'") case True hence *: "of_circline (of_ocircline ?H') = ?H'" using of_circline_of_ocircline_pos_oriented by blast from ounit show ?thesis proof assume **: "ounit_circle = moebius_ocircline M ounit_circle" show ?thesis using ++ unfolding perpendicular_def by (simp, subst moebius_circline_ocircline, subst *, subst **) simp next assume **: "ounit_circle = moebius_ocircline M (opposite_ocircline ounit_circle)" show ?thesis using ++ unfolding perpendicular_def by (simp, subst moebius_circline_ocircline, subst *, subst **) simp qed next case False hence *: "of_circline (of_ocircline ?H') = opposite_ocircline ?H'" by (metis of_circline_of_ocircline pos_oriented_of_circline) from ounit show ?thesis proof assume **: "ounit_circle = moebius_ocircline M ounit_circle" show ?thesis using ++ unfolding perpendicular_def by (simp, subst moebius_circline_ocircline, subst *, subst **) simp next assume **: "ounit_circle = moebius_ocircline M (opposite_ocircline ounit_circle)" show ?thesis using ++ unfolding perpendicular_def by (simp, subst moebius_circline_ocircline, subst *, subst **) simp qed qed qed simp lemma unit_circle_fix_preserve_is_poincare_line_iff [simp]: assumes "unit_circle_fix M" shows "is_poincare_line (moebius_circline M H) \ is_poincare_line H" using assms using unit_circle_fix_preserve_is_poincare_line[of M H] using unit_circle_fix_preserve_is_poincare_line[of "moebius_inv M" "moebius_circline M H"] by (auto simp del: unit_circle_fix_preserve_is_poincare_line) text\Since h-lines are preserved by transformations that fix the unit circle, so is collinearity.\ lemma unit_disc_fix_preserve_poincare_collinear [simp]: assumes "unit_circle_fix M" "poincare_collinear A" shows "poincare_collinear (moebius_pt M ` A)" using assms unfolding poincare_collinear_def by (auto, rule_tac x="moebius_circline M p" in exI, auto) lemma unit_disc_fix_preserve_poincare_collinear_iff [simp]: assumes "unit_circle_fix M" shows "poincare_collinear (moebius_pt M ` A) \ poincare_collinear A" using assms using unit_disc_fix_preserve_poincare_collinear[of M A] using unit_disc_fix_preserve_poincare_collinear[of "moebius_inv M" "moebius_pt M ` A"] by (auto simp del: unit_disc_fix_preserve_poincare_collinear) lemma unit_disc_fix_preserve_poincare_collinear3 [simp]: assumes "unit_disc_fix M" shows "poincare_collinear {moebius_pt M u, moebius_pt M v, moebius_pt M w} \ poincare_collinear {u, v, w}" using assms unit_disc_fix_preserve_poincare_collinear_iff[of M "{u, v, w}"] by simp text\Conjugation is also an h-isometry and it preserves h-lines.\ lemma is_poincare_line_conjugate_circline [simp]: assumes "is_poincare_line H" shows "is_poincare_line (conjugate_circline H)" using assms by (transfer, transfer, auto simp add: mat_cnj_def hermitean_def mat_adj_def) lemma is_poincare_line_conjugate_circline_iff [simp]: shows "is_poincare_line (conjugate_circline H) \ is_poincare_line H" using is_poincare_line_conjugate_circline[of "conjugate_circline H"] by auto text\Since h-lines are preserved by conjugation, so is collinearity.\ lemma conjugate_preserve_poincare_collinear [simp]: assumes "poincare_collinear A" shows "poincare_collinear (conjugate ` A)" using assms unfolding poincare_collinear_def by auto (rule_tac x="conjugate_circline p" in exI, auto) lemma conjugate_conjugate [simp]: "conjugate ` conjugate ` A = A" by (auto simp add: image_iff) lemma conjugate_preserve_poincare_collinear_iff [simp]: shows "poincare_collinear (conjugate ` A) \ poincare_collinear A" using conjugate_preserve_poincare_collinear[of "A"] using conjugate_preserve_poincare_collinear[of "conjugate ` A"] by (auto simp del: conjugate_preserve_poincare_collinear) (* ------------------------------------------------------------------ *) subsubsection\Mapping h-lines to x-axis\ (* ------------------------------------------------------------------ *) text\Each h-line in the Poincar\'e model can be mapped onto the x-axis (by a unit-disc preserving Möbius transformation).\ lemma ex_unit_disc_fix_is_poincare_line_to_x_axis: assumes "is_poincare_line l" shows "\ M. unit_disc_fix M \ moebius_circline M l = x_axis" proof- from assms obtain u v where "u \ v" "u \ unit_disc" "v \ unit_disc" and "{u, v} \ circline_set l" using ex_is_poincare_line_points by blast then 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[of u v] by auto moreover hence "{0\<^sub>h, moebius_pt M v} \ circline_set x_axis" unfolding positive_x_axis_def by auto moreover have "moebius_pt M v \ 0\<^sub>h" using \u \ v\ * by (metis moebius_pt_neq_I) moreover have "moebius_pt M v \ \\<^sub>h" using \unit_disc_fix M\ \v \ unit_disc\ using unit_disc_fix_discI by fastforce ultimately show ?thesis using \is_poincare_line l\ \{u, v} \ circline_set l\ \unit_disc_fix M\ using is_poincare_line_0_real_is_x_axis[of "moebius_circline M l" "moebius_pt M v"] by (rule_tac x="M" in exI, force) qed text \When proving facts about h-lines, without loss of generality it can be assumed that h-line is the x-axis (if the property being proved is invariant under Möbius transformations that fix the unit disc).\ lemma wlog_line_x_axis: assumes is_line: "is_poincare_line H" assumes x_axis: "P x_axis" assumes preserves: "\ M. \unit_disc_fix M; P (moebius_circline M H)\ \ P H" shows "P H" using assms using ex_unit_disc_fix_is_poincare_line_to_x_axis[of H] by auto (* ------------------------------------------------------------------ *) subsection\Construction of the h-line between the two given points\ (* ------------------------------------------------------------------ *) text\Next we show how to construct the (unique) h-line between the two given points in the Poincar\'e model\ text\ Geometrically, h-line can be constructed by finding the inverse point of one of the two points and by constructing the circle (or line) trough it and the two given points. Algebraically, for two given points $u$ and $v$ in $\mathbb{C}$, the h-line matrix coefficients can be $A = i\cdot(u\overline{v}-v\overline{u})$ and $B = i\cdot(v(|u|^2+1) - u(|v|^2+1))$. We need to extend this to homogenous coordinates. There are several degenerate cases. - If $\{z, w\} = \{0_h, \infty_h\}$ then there is no unique h-line (any line trough zero is an h-line). - If z and w are mutually inverse, then the construction fails (both geometric and algebraic). - If z and w are different points on the unit circle, then the standard construction fails (only geometric). - None of this problematic cases occur when z and w are inside the unit disc. We express the construction algebraically, and construct the Hermitean circline matrix for the two points given in homogenous coordinates. It works correctly in all cases except when the two points are the same or are mutually inverse. \ definition mk_poincare_line_cmat :: "real \ complex \ complex_mat" where [simp]: "mk_poincare_line_cmat A B = (cor A, B, cnj B, cor A)" lemma mk_poincare_line_cmat_zero_iff: "mk_poincare_line_cmat A B = mat_zero \ A = 0 \ B = 0" by auto lemma mk_poincare_line_cmat_hermitean [simp]: "hermitean (mk_poincare_line_cmat A B)" by simp lemma mk_poincare_line_cmat_scale: "cor k *\<^sub>s\<^sub>m mk_poincare_line_cmat A B = mk_poincare_line_cmat (k * A) (k * B)" by simp definition poincare_line_cvec_cmat :: "complex_vec \ complex_vec \ complex_mat" where [simp]: "poincare_line_cvec_cmat z w = (let (z1, z2) = z; (w1, w2) = w; nom = w1*cnj w2*(z1*cnj z1 + z2*cnj z2) - z1*cnj z2*(w1*cnj w1 + w2*cnj w2); den = z1*cnj z2*cnj w1*w2 - w1*cnj w2*cnj z1*z2 in if den \ 0 then mk_poincare_line_cmat (Re(\*den)) (\*nom) else if z1*cnj z2 \ 0 then mk_poincare_line_cmat 0 (\*z1*cnj z2) else if w1*cnj w2 \ 0 then mk_poincare_line_cmat 0 (\*w1*cnj w2) else mk_poincare_line_cmat 0 \)" lemma poincare_line_cvec_cmat_AeqD: assumes "poincare_line_cvec_cmat z w = (A, B, C, D)" shows "A = D" using assms by (cases z, cases w) (auto split: if_split_asm) lemma poincare_line_cvec_cmat_hermitean [simp]: shows "hermitean (poincare_line_cvec_cmat z w)" by (cases z, cases w) (auto split: if_split_asm simp del: mk_poincare_line_cmat_def) lemma poincare_line_cvec_cmat_nonzero [simp]: assumes "z \ vec_zero" "w \ vec_zero" shows "poincare_line_cvec_cmat z w \ mat_zero" proof- obtain z1 z2 w1 w2 where *: "z = (z1, z2)" "w = (w1, w2)" by (cases z, cases w, auto) let ?den = "z1*cnj z2*cnj w1*w2 - w1*cnj w2*cnj z1*z2" show ?thesis proof (cases "?den \ 0") case True have "is_real (\ * ?den)" using eq_cnj_iff_real[of "\ *?den"] by (simp add: field_simps) hence "Re (\ * ?den) \ 0" using \?den \ 0\ by (metis complex_i_not_zero complex_surj mult_eq_0_iff zero_complex.code) thus ?thesis using * \?den \ 0\ by (simp del: mk_poincare_line_cmat_def mat_zero_def add: mk_poincare_line_cmat_zero_iff) next case False thus ?thesis using * by (simp del: mk_poincare_line_cmat_def mat_zero_def add: mk_poincare_line_cmat_zero_iff) qed qed lift_definition poincare_line_hcoords_clmat :: "complex_homo_coords \ complex_homo_coords \ circline_mat" is poincare_line_cvec_cmat using poincare_line_cvec_cmat_hermitean poincare_line_cvec_cmat_nonzero by simp lift_definition poincare_line :: "complex_homo \ complex_homo \ circline" is poincare_line_hcoords_clmat proof transfer fix za zb wa wb assume "za \ vec_zero" "zb \ vec_zero" "wa \ vec_zero" "wb \ vec_zero" assume "za \\<^sub>v zb" "wa \\<^sub>v wb" obtain za1 za2 zb1 zb2 wa1 wa2 wb1 wb2 where *: "(za1, za2) = za" "(zb1, zb2) = zb" "(wa1, wa2) = wa" "(wb1, wb2) = wb" by (cases za, cases zb, cases wa, cases wb, auto) obtain kz kw where **: "kz \ 0" "kw \ 0" "zb1 = kz * za1" "zb2 = kz * za2" "wb1 = kw * wa1" "wb2 = kw * wa2" using \za \\<^sub>v zb\ \wa \\<^sub>v wb\ *[symmetric] by auto let ?nom = "\ z1 z2 w1 w2. w1*cnj w2*(z1*cnj z1 + z2*cnj z2) - z1*cnj z2*(w1*cnj w1 + w2*cnj w2)" let ?den = "\ z1 z2 w1 w2. z1*cnj z2*cnj w1*w2 - w1*cnj w2*cnj z1*z2" show "circline_eq_cmat (poincare_line_cvec_cmat za wa) (poincare_line_cvec_cmat zb wb)" proof- have "\k. k \ 0 \ poincare_line_cvec_cmat (zb1, zb2) (wb1, wb2) = cor k *\<^sub>s\<^sub>m poincare_line_cvec_cmat (za1, za2) (wa1, wa2)" proof (cases "?den za1 za2 wa1 wa2 \ 0") case True hence "?den zb1 zb2 wb1 wb2 \ 0" using ** by (simp add: field_simps) let ?k = "kz * cnj kz * kw * cnj kw" have "?k \ 0" using ** by simp have "is_real ?k" using eq_cnj_iff_real[of ?k] by auto have "cor (Re ?k) = ?k" using \is_real ?k\ using complex_of_real_Re by blast have "Re ?k \ 0" using \?k \ 0\ \cor (Re ?k) = ?k\ by (metis of_real_0) have arg1: "Re (\ * ?den zb1 zb2 wb1 wb2) = Re ?k * Re (\ * ?den za1 za2 wa1 wa2)" apply (subst **)+ apply (subst Re_mult_real[symmetric, OF \is_real ?k\]) apply (rule arg_cong[where f=Re]) apply (simp add: field_simps) done have arg2: "\ * ?nom zb1 zb2 wb1 wb2 = ?k * \ * ?nom za1 za2 wa1 wa2" using ** by (simp add: field_simps) have "mk_poincare_line_cmat (Re (\*?den zb1 zb2 wb1 wb2)) (\*?nom zb1 zb2 wb1 wb2) = cor (Re ?k) *\<^sub>s\<^sub>m mk_poincare_line_cmat (Re (\*?den za1 za2 wa1 wa2)) (\*?nom za1 za2 wa1 wa2)" using \cor (Re ?k) = ?k\ \is_real ?k\ apply (subst mk_poincare_line_cmat_scale) apply (subst arg1, subst arg2) apply (subst \cor (Re ?k) = ?k\)+ apply simp done thus ?thesis using \?den za1 za2 wa1 wa2 \ 0\ \?den zb1 zb2 wb1 wb2 \ 0\ using \Re ?k \ 0\ \cor (Re ?k) = ?k\ by (rule_tac x="Re ?k" in exI, simp) next case False hence "?den zb1 zb2 wb1 wb2 = 0" using ** by (simp add: field_simps) show ?thesis proof (cases "za1*cnj za2 \ 0") case True hence "zb1*cnj zb2 \ 0" using ** by (simp add: field_simps) let ?k = "kz * cnj kz" have "?k \ 0" "is_real ?k" using ** using eq_cnj_iff_real[of ?k] by auto thus ?thesis using \za1 * cnj za2 \ 0\ \zb1 * cnj zb2 \ 0\ using \\ (?den za1 za2 wa1 wa2 \ 0)\ \?den zb1 zb2 wb1 wb2 = 0\ ** by (rule_tac x="Re (kz * cnj kz)" in exI, auto simp add: complex.expand) next case False hence "zb1 * cnj zb2 = 0" using ** by (simp add: field_simps) show ?thesis proof (cases "wa1 * cnj wa2 \ 0") case True hence "wb1*cnj wb2 \ 0" using ** by (simp add: field_simps) let ?k = "kw * cnj kw" have "?k \ 0" "is_real ?k" using ** using eq_cnj_iff_real[of ?k] by auto thus ?thesis using \\ (za1 * cnj za2 \ 0)\ using \wa1 * cnj wa2 \ 0\ \wb1 * cnj wb2 \ 0\ using \\ (?den za1 za2 wa1 wa2 \ 0)\ \?den zb1 zb2 wb1 wb2 = 0\ ** by (rule_tac x="Re (kw * cnj kw)" in exI) (auto simp add: complex.expand) next case False hence "wb1 * cnj wb2 = 0" using ** by (simp add: field_simps) thus ?thesis using \\ (za1 * cnj za2 \ 0)\ \zb1 * cnj zb2 = 0\ using \\ (wa1 * cnj wa2 \ 0)\ \wb1 * cnj wb2 = 0\ using \\ (?den za1 za2 wa1 wa2 \ 0)\ \?den zb1 zb2 wb1 wb2 = 0\ ** by simp qed qed qed thus ?thesis using *[symmetric] by simp qed qed subsubsection \Correctness of the construction\ text\For finite points, our definition matches the classic algebraic definition for points in $\mathbb{C}$ (given in ordinary, not homogenous coordinates).\ lemma poincare_line_non_homogenous: assumes "u \ \\<^sub>h" "v \ \\<^sub>h" "u \ v" "u \ inversion v" shows "let u' = to_complex u; v' = to_complex v; A = \ * (u' * cnj v' - v' * cnj u'); B = \ * (v' * ((cmod u')\<^sup>2 + 1) - u' * ((cmod v')\<^sup>2 + 1)) in poincare_line u v = mk_circline A B (cnj B) A" using assms unfolding unit_disc_def disc_def inversion_def apply (simp add: Let_def) proof (transfer, transfer, safe) fix u1 u2 v1 v2 assume uv: "(u1, u2) \ vec_zero" "(v1, v2) \ vec_zero" "\ (u1, u2) \\<^sub>v \\<^sub>v" "\ (v1, v2) \\<^sub>v \\<^sub>v" "\ (u1, u2) \\<^sub>v (v1, v2)" "\ (u1, u2) \\<^sub>v conjugate_cvec (reciprocal_cvec (v1, v2))" let ?u = "to_complex_cvec (u1, u2)" and ?v = "to_complex_cvec (v1, v2)" let ?A = "\ * (?u * cnj ?v - ?v * cnj ?u)" let ?B = "\ * (?v * ((cor (cmod ?u))\<^sup>2 + 1) - ?u * ((cor (cmod ?v))\<^sup>2 + 1))" let ?C = "- (\ * (cnj ?v * ((cor (cmod ?u))\<^sup>2 + 1) - cnj ?u * ((cor (cmod ?v))\<^sup>2 + 1)))" let ?D = ?A let ?H = "(?A, ?B, ?C, ?D)" let ?den = "u1 * cnj u2 * cnj v1 * v2 - v1 * cnj v2 * cnj u1 * u2" have "u2 \ 0" "v2 \ 0" using uv using inf_cvec_z2_zero_iff by blast+ have "\ (u1, u2) \\<^sub>v (cnj v2, cnj v1)" using uv(6) by (simp add: vec_cnj_def) moreover have "(cnj v2, cnj v1) \ vec_zero" using uv(2) by auto ultimately have *: "u1 * cnj v1 \ u2 * cnj v2" "u1 * v2 \ u2 * v1" using uv(5) uv(1) uv(2) \u2 \ 0\ \v2 \ 0\ using complex_cvec_eq_mix by blast+ show "circline_eq_cmat (poincare_line_cvec_cmat (u1, u2) (v1, v2)) (mk_circline_cmat ?A ?B ?C ?D)" proof (cases "?den \ 0") case True let ?nom = "v1 * cnj v2 * (u1 * cnj u1 + u2 * cnj u2) - u1 * cnj u2 * (v1 * cnj v1 + v2 * cnj v2)" let ?H' = "mk_poincare_line_cmat (Re (\ * ?den)) (\ * ?nom)" have "circline_eq_cmat ?H ?H'" proof- let ?k = "(u2 * cnj v2) * (v2 * cnj u2)" have "is_real ?k" using eq_cnj_iff_real by fastforce hence "cor (Re ?k) = ?k" using complex_of_real_Re by blast have "Re (\ * ?den) = Re ?k * ?A" proof- have "?A = cnj ?A" by (simp add: field_simps) hence "is_real ?A" using eq_cnj_iff_real by fastforce moreover have "\ * ?den = cnj (\ * ?den)" by (simp add: field_simps) hence "is_real (\ * ?den)" using eq_cnj_iff_real by fastforce hence "cor (Re (\ * ?den)) = \ * ?den" using complex_of_real_Re by blast ultimately show ?thesis using \cor (Re ?k) = ?k\ by (simp add: field_simps) qed moreover have "\ * ?nom = Re ?k * ?B" using \cor (Re ?k) = ?k\ \u2 \ 0\ \v2 \ 0\ complex_mult_cnj_cmod[symmetric] by (auto simp add: field_simps) moreover have "?k \ 0" using \u2 \ 0\ \v2 \ 0\ by simp hence "Re ?k \ 0" using \is_real ?k\ by (metis \cor (Re ?k) = ?k\ of_real_0) ultimately show ?thesis by simp (rule_tac x="Re ?k" in exI, simp add: mult.commute) qed moreover have "poincare_line_cvec_cmat (u1, u2) (v1, v2) = ?H'" using \?den \ 0\ unfolding poincare_line_cvec_cmat_def by (simp add: Let_def) moreover hence "hermitean ?H' \ ?H' \ mat_zero" by (metis mk_poincare_line_cmat_hermitean poincare_line_cvec_cmat_nonzero uv(1) uv(2)) hence "hermitean ?H \ ?H \ mat_zero" using \circline_eq_cmat ?H ?H'\ using circline_eq_cmat_hermitean_nonzero[of ?H' ?H] symp_circline_eq_cmat unfolding symp_def by metis hence "mk_circline_cmat ?A ?B ?C ?D = ?H" by simp ultimately have "circline_eq_cmat (mk_circline_cmat ?A ?B ?C ?D) (poincare_line_cvec_cmat (u1, u2) (v1, v2))" by simp thus ?thesis using symp_circline_eq_cmat unfolding symp_def by blast next case False let ?d = "v1 * (u1 * cnj u1 / (u2 * cnj u2) + 1) / v2 - u1 * (v1 * cnj v1 / (v2 * cnj v2) + 1) / u2" let ?cd = "cnj v1 * (u1 * cnj u1 / (u2 * cnj u2) + 1) / cnj v2 - cnj u1 * (v1 * cnj v1 / (v2 * cnj v2) + 1) / cnj u2" have "cnj ?d = ?cd" by (simp add: mult.commute) let ?d1 = "(v1 / v2) * (cnj u1 / cnj u2) - 1" let ?d2 = "u1 / u2 - v1 / v2" have **: "?d = ?d1 * ?d2" using \\ ?den \ 0\ \u2 \ 0\ \v2 \ 0\ by(simp add: field_simps) hence "?d \ 0" using \\ ?den \ 0\ \u2 \ 0\ \v2 \ 0\ * by auto (simp add: field_simps)+ have "is_real ?d1" proof- have "cnj ?d1 = ?d1" using \\ ?den \ 0\ \u2 \ 0\ \v2 \ 0\ * by (simp add: field_simps) thus ?thesis using eq_cnj_iff_real by blast qed show ?thesis proof (cases "u1 * cnj u2 \ 0") case True let ?nom = "u1 * cnj u2" let ?H' = "mk_poincare_line_cmat 0 (\ * ?nom)" have "circline_eq_cmat ?H ?H'" proof- let ?k = "(u1 * cnj u2) / ?d" have "is_real ?k" proof- have "is_real ((u1 * cnj u2) / ?d2)" proof- let ?rhs = "(u2 * cnj u2) / (1 - (v1*u2)/(u1*v2))" have 1: "(u1 * cnj u2) / ?d2 = ?rhs" using \\ ?den \ 0\ \u2 \ 0\ \v2 \ 0\ * \u1 * cnj u2 \ 0\ by (simp add: field_simps) moreover have "cnj ?rhs = ?rhs" proof- have "cnj (1 - v1 * u2 / (u1 * v2)) = 1 - v1 * u2 / (u1 * v2)" using \\ ?den \ 0\ \u2 \ 0\ \v2 \ 0\ * \u1 * cnj u2 \ 0\ by (simp add: field_simps) moreover have "cnj (u2 * cnj u2) = u2 * cnj u2" by simp ultimately show ?thesis by simp qed ultimately show ?thesis using eq_cnj_iff_real by fastforce qed thus ?thesis using ** \is_real ?d1\ by (metis complex_cnj_divide divide_divide_eq_left' eq_cnj_iff_real) qed have "?k \ 0" using \?d \ 0\ \u1 * cnj u2 \ 0\ by simp have "cnj ?k = ?k" using \is_real ?k\ using eq_cnj_iff_real by blast have "Re ?k \ 0" using \?k \ 0\ \is_real ?k\ by (metis complex.expand zero_complex.simps(1) zero_complex.simps(2)) have "u1 * cnj u2 = ?k * ?d" using \?d \ 0\ by simp moreover hence "cnj u1 * u2 = cnj ?k * cnj ?d" by (metis complex_cnj_cnj complex_cnj_mult) hence "cnj u1 * u2 = ?k * ?cd" using \cnj ?k = ?k\ \cnj ?d = ?cd\ by metis ultimately show ?thesis using \~ ?den \ 0\ \u1 * cnj u2 \ 0\ \u2 \ 0\ \v2 \ 0\ \Re ?k \ 0\ \is_real ?k\ \?d \ 0\ using complex_mult_cnj_cmod[symmetric, of u1] using complex_mult_cnj_cmod[symmetric, of v1] using complex_mult_cnj_cmod[symmetric, of u2] using complex_mult_cnj_cmod[symmetric, of v2] apply (simp add: power_divide norm_mult norm_divide) apply (rule_tac x="Re ?k" in exI) apply simp apply (simp add: field_simps) done qed moreover have "poincare_line_cvec_cmat (u1, u2) (v1, v2) = ?H'" using \\ ?den \ 0\ \u1 * cnj u2 \ 0\ unfolding poincare_line_cvec_cmat_def by (simp add: Let_def) moreover hence "hermitean ?H' \ ?H' \ mat_zero" by (metis mk_poincare_line_cmat_hermitean poincare_line_cvec_cmat_nonzero uv(1) uv(2)) hence "hermitean ?H \ ?H \ mat_zero" using \circline_eq_cmat ?H ?H'\ using circline_eq_cmat_hermitean_nonzero[of ?H' ?H] symp_circline_eq_cmat unfolding symp_def by metis hence "mk_circline_cmat ?A ?B ?C ?D = ?H" by simp ultimately have "circline_eq_cmat (mk_circline_cmat ?A ?B ?C ?D) (poincare_line_cvec_cmat (u1, u2) (v1, v2))" by simp thus ?thesis using symp_circline_eq_cmat unfolding symp_def by blast next case False show ?thesis proof (cases "v1 * cnj v2 \ 0") case True let ?nom = "v1 * cnj v2" let ?H' = "mk_poincare_line_cmat 0 (\ * ?nom)" have "circline_eq_cmat ?H ?H'" proof- let ?k = "(v1 * cnj v2) / ?d" have "is_real ?k" proof- have "is_real ((v1 * cnj v2) / ?d2)" proof- let ?rhs = "(v2 * cnj v2) / ((u1*v2)/(u2*v1) - 1)" have 1: "(v1 * cnj v2) / ?d2 = ?rhs" using \\ ?den \ 0\ \u2 \ 0\ \v2 \ 0\ * \v1 * cnj v2 \ 0\ by (simp add: field_simps) moreover have "cnj ?rhs = ?rhs" proof- have "cnj (u1 * v2 / (u2 * v1) - 1) = u1 * v2 / (u2 * v1) - 1" using \\ ?den \ 0\ \u2 \ 0\ \v2 \ 0\ * \v1 * cnj v2 \ 0\ by (simp add: field_simps) moreover have "cnj (v2 * cnj v2) = v2 * cnj v2" by simp ultimately show ?thesis by simp qed ultimately show ?thesis using eq_cnj_iff_real by fastforce qed thus ?thesis using ** \is_real ?d1\ by (metis complex_cnj_divide divide_divide_eq_left' eq_cnj_iff_real) qed have "?k \ 0" using \?d \ 0\ \v1 * cnj v2 \ 0\ by simp have "cnj ?k = ?k" using \is_real ?k\ using eq_cnj_iff_real by blast have "Re ?k \ 0" using \?k \ 0\ \is_real ?k\ by (metis complex.expand zero_complex.simps(1) zero_complex.simps(2)) have "v1 * cnj v2 = ?k * ?d" using \?d \ 0\ by simp moreover hence "cnj v1 * v2 = cnj ?k * cnj ?d" by (metis complex_cnj_cnj complex_cnj_mult) hence "cnj v1 * v2 = ?k * ?cd" using \cnj ?k = ?k\ \cnj ?d = ?cd\ by metis ultimately show ?thesis using \~ ?den \ 0\ \v1 * cnj v2 \ 0\ \u2 \ 0\ \v2 \ 0\ \Re ?k \ 0\ \is_real ?k\ \?d \ 0\ using complex_mult_cnj_cmod[symmetric, of u1] using complex_mult_cnj_cmod[symmetric, of v1] using complex_mult_cnj_cmod[symmetric, of u2] using complex_mult_cnj_cmod[symmetric, of v2] apply (simp add: power_divide norm_mult norm_divide) apply (rule_tac x="Re ?k" in exI) apply simp apply (simp add: field_simps) done qed moreover have "poincare_line_cvec_cmat (u1, u2) (v1, v2) = ?H'" using \\ ?den \ 0\ \\ u1 * cnj u2 \ 0\ \v1 * cnj v2 \ 0\ unfolding poincare_line_cvec_cmat_def by (simp add: Let_def) moreover hence "hermitean ?H' \ ?H' \ mat_zero" by (metis mk_poincare_line_cmat_hermitean poincare_line_cvec_cmat_nonzero uv(1) uv(2)) hence "hermitean ?H \ ?H \ mat_zero" using \circline_eq_cmat ?H ?H'\ using circline_eq_cmat_hermitean_nonzero[of ?H' ?H] symp_circline_eq_cmat unfolding symp_def by metis hence "mk_circline_cmat ?A ?B ?C ?D = ?H" by simp ultimately have "circline_eq_cmat (mk_circline_cmat ?A ?B ?C ?D) (poincare_line_cvec_cmat (u1, u2) (v1, v2))" by simp thus ?thesis using symp_circline_eq_cmat unfolding symp_def by blast next case False hence False using \\ ?den \ 0\ \\ u1 * cnj u2 \ 0\ uv by (simp add: \u2 \ 0\ \v2 \ 0\) thus ?thesis by simp qed qed qed qed text\Our construction (in homogenous coordinates) always yields an h-line that contain two starting points (this also holds for all degenerate cases except when points are the same).\ lemma poincare_line [simp]: assumes "z \ w" shows "on_circline (poincare_line z w) z" "on_circline (poincare_line z w) w" proof- have "on_circline (poincare_line z w) z \ on_circline (poincare_line z w) w" using assms proof (transfer, transfer) fix z w assume vz: "z \ vec_zero" "w \ vec_zero" obtain z1 z2 w1 w2 where zw: "(z1, z2) = z" "(w1, w2) = w" by (cases z, cases w, auto) let ?den = "z1*cnj z2*cnj w1*w2 - w1*cnj w2*cnj z1*z2" have *: "cor (Re (\ * ?den)) = \ * ?den" proof- have "cnj ?den = -?den" by auto hence "is_imag ?den" using eq_minus_cnj_iff_imag[of ?den] by simp thus ?thesis using complex_of_real_Re[of "\ * ?den"] by simp qed show "on_circline_cmat_cvec (poincare_line_cvec_cmat z w) z \ on_circline_cmat_cvec (poincare_line_cvec_cmat z w) w" unfolding poincare_line_cvec_cmat_def mk_poincare_line_cmat_def apply (subst zw[symmetric])+ unfolding Let_def prod.case apply (subst *)+ by (auto simp add: vec_cnj_def field_simps) qed thus "on_circline (poincare_line z w) z" "on_circline (poincare_line z w) w" by auto qed lemma poincare_line_circline_set [simp]: assumes "z \ w" shows "z \ circline_set (poincare_line z w)" "w \ circline_set (poincare_line z w)" using assms by (auto simp add: circline_set_def) text\When the points are different, the constructed line matrix always has a negative determinant\ lemma poincare_line_type: assumes "z \ w" shows "circline_type (poincare_line z w) = -1" proof- have "\ a b. a \ b \ {a, b} \ circline_set (poincare_line z w)" using poincare_line[of z w] assms unfolding circline_set_def by (rule_tac x=z in exI, rule_tac x=w in exI, simp) thus ?thesis using circline_type[of "poincare_line z w"] using circline_type_pos_card_eq0[of "poincare_line z w"] using circline_type_zero_card_eq1[of "poincare_line z w"] by auto qed text\The constructed line is an h-line in the Poincar\'e model (in all cases when the two points are different)\ lemma is_poincare_line_poincare_line [simp]: assumes "z \ w" shows "is_poincare_line (poincare_line z w)" using poincare_line_type[of z w, OF assms] proof (transfer, transfer) fix z w assume vz: "z \ vec_zero" "w \ vec_zero" obtain A B C D where *: "poincare_line_cvec_cmat z w = (A, B, C, D)" by (cases "poincare_line_cvec_cmat z w") auto assume "circline_type_cmat (poincare_line_cvec_cmat z w) = - 1" thus "is_poincare_line_cmat (poincare_line_cvec_cmat z w)" using vz * using poincare_line_cvec_cmat_hermitean[of z w] using poincare_line_cvec_cmat_nonzero[of z w] using poincare_line_cvec_cmat_AeqD[of z w A B C D] using hermitean_elems[of A B C D] using cmod_power2[of D] cmod_power2[of C] unfolding is_poincare_line_cmat_def by (simp del: poincare_line_cvec_cmat_def add: sgn_1_neg power2_eq_square) qed text \When the points are different, the constructed h-line between two points also contains their inverses\ lemma poincare_line_inversion: assumes "z \ w" shows "on_circline (poincare_line z w) (inversion z)" "on_circline (poincare_line z w) (inversion w)" using assms using is_poincare_line_poincare_line[OF \z \ w\] using is_poincare_line_inverse_point unfolding circline_set_def by auto text \When the points are different, the onstructed h-line between two points contains the inverse of its every point\ lemma poincare_line_inversion_full: assumes "u \ v" assumes "on_circline (poincare_line u v) x" shows "on_circline (poincare_line u v) (inversion x)" using is_poincare_line_inverse_point[of "poincare_line u v" x] using is_poincare_line_poincare_line[OF \u \ v\] assms unfolding circline_set_def by simp subsubsection \Existence of h-lines\ text\There is an h-line trough every point in the Poincar\'e model\ lemma ex_poincare_line_one_point: shows "\ l. is_poincare_line l \ z \ circline_set l" proof (cases "z = 0\<^sub>h") case True thus ?thesis by (rule_tac x="x_axis" in exI) simp next case False thus ?thesis by (rule_tac x="poincare_line 0\<^sub>h z" in exI) auto qed lemma poincare_collinear_singleton [simp]: assumes "u \ unit_disc" shows "poincare_collinear {u}" using assms using ex_poincare_line_one_point[of u] by (auto simp add: poincare_collinear_def) text\There is an h-line trough every two points in the Poincar\'e model\ lemma ex_poincare_line_two_points: assumes "z \ w" shows "\ l. is_poincare_line l \ z \ circline_set l \ w \ circline_set l" using assms by (rule_tac x="poincare_line z w" in exI, simp) lemma poincare_collinear_doubleton [simp]: assumes "u \ unit_disc" "v \ unit_disc" shows "poincare_collinear {u, v}" using assms using ex_poincare_line_one_point[of u] using ex_poincare_line_two_points[of u v] by (cases "u = v") (simp_all add: poincare_collinear_def) subsubsection \Uniqueness of h-lines\ text \The only h-line between two points is the one obtained by the line-construction.\ text \First we show this only for two different points inside the disc.\ lemma unique_poincare_line: assumes in_disc: "u \ v" "u \ unit_disc" "v \ unit_disc" assumes on_l: "u \ circline_set l" "v \ circline_set l" "is_poincare_line l" shows "l = poincare_line u v" using assms using unique_is_poincare_line[of u v l "poincare_line u v"] unfolding circline_set_def by auto text\The assumption that the points are inside the disc can be relaxed.\ lemma unique_poincare_line_general: assumes in_disc: "u \ v" "u \ inversion v" assumes on_l: "u \ circline_set l" "v \ circline_set l" "is_poincare_line l" shows "l = poincare_line u v" using assms using unique_is_poincare_line_general[of u v l "poincare_line u v"] unfolding circline_set_def by auto text\The explicit line construction enables us to prove that there exists a unique h-line through any given two h-points (uniqueness part was already shown earlier).\ text \First we show this only for two different points inside the disc.\ lemma ex1_poincare_line: assumes "u \ v" "u \ unit_disc" "v \ unit_disc" shows "\! l. is_poincare_line l \ u \ circline_set l \ v \ circline_set l" proof (rule ex1I) let ?l = "poincare_line u v" show "is_poincare_line ?l \ u \ circline_set ?l \ v \ circline_set ?l" using assms unfolding circline_set_def by auto next fix l assume "is_poincare_line l \ u \ circline_set l \ v \ circline_set l" thus "l = poincare_line u v" using unique_poincare_line assms by auto qed text \The assumption that the points are in the disc can be relaxed.\ lemma ex1_poincare_line_general: assumes "u \ v" "u \ inversion v" shows "\! l. is_poincare_line l \ u \ circline_set l \ v \ circline_set l" proof (rule ex1I) let ?l = "poincare_line u v" show "is_poincare_line ?l \ u \ circline_set ?l \ v \ circline_set ?l" using assms unfolding circline_set_def by auto next fix l assume "is_poincare_line l \ u \ circline_set l \ v \ circline_set l" thus "l = poincare_line u v" using unique_poincare_line_general assms by auto qed subsubsection \Some consequences of line uniqueness\ text\H-line $uv$ is the same as the h-line $vu$.\ lemma poincare_line_sym: assumes "u \ unit_disc" "v \ unit_disc" "u \ v" shows "poincare_line u v = poincare_line v u" using assms using unique_poincare_line[of u v "poincare_line v u"] by simp lemma poincare_line_sym_general: assumes "u \ v" "u \ inversion v" shows "poincare_line u v = poincare_line v u" using assms using unique_poincare_line_general[of u v "poincare_line v u"] by simp text\Each h-line is the h-line constructed out of its two arbitrary different points.\ lemma ex_poincare_line_points: assumes "is_poincare_line H" shows "\ u v. u \ unit_disc \ v \ unit_disc \ u \ v \ H = poincare_line u v" using assms using ex_is_poincare_line_points using unique_poincare_line[where l=H] by fastforce text\If an h-line contains two different points on x-axis/y-axis then it is the x-axis/y-axis.\ lemma poincare_line_0_real_is_x_axis: assumes "x \ circline_set x_axis" "x \ 0\<^sub>h" "x \ \\<^sub>h" shows "poincare_line 0\<^sub>h x = x_axis" using assms using is_poincare_line_0_real_is_x_axis[of "poincare_line 0\<^sub>h x" x] by auto lemma poincare_line_0_imag_is_y_axis: assumes "y \ circline_set y_axis" "y \ 0\<^sub>h" "y \ \\<^sub>h" shows "poincare_line 0\<^sub>h y = y_axis" using assms using is_poincare_line_0_imag_is_y_axis[of "poincare_line 0\<^sub>h y" y] by auto lemma poincare_line_x_axis: assumes "x \ unit_disc" "y \ unit_disc" "x \ circline_set x_axis" "y \ circline_set x_axis" "x \ y" shows "poincare_line x y = x_axis" using assms using unique_poincare_line by auto lemma poincare_line_minus_one_one [simp]: shows "poincare_line (of_complex (-1)) (of_complex 1) = x_axis" proof- have "0\<^sub>h \ circline_set (poincare_line (of_complex (-1)) (of_complex 1))" unfolding circline_set_def by simp (transfer, transfer, simp add: vec_cnj_def) hence "poincare_line 0\<^sub>h (of_complex 1) = poincare_line (of_complex (-1)) (of_complex 1)" by (metis is_poincare_line_poincare_line is_poincare_line_trough_zero_trough_infty not_zero_on_unit_circle of_complex_inj of_complex_one one_neq_neg_one one_on_unit_circle poincare_line_0_real_is_x_axis poincare_line_circline_set(2) reciprocal_involution reciprocal_one reciprocal_zero unique_circline_01inf') thus ?thesis using poincare_line_0_real_is_x_axis[of "of_complex 1"] by auto qed subsubsection \Transformations of constructed lines\ text\Unit dics preserving Möbius transformations preserve the h-line construction\ lemma unit_disc_fix_preserve_poincare_line [simp]: assumes "unit_disc_fix M" "u \ unit_disc" "v \ unit_disc" "u \ v" shows "poincare_line (moebius_pt M u) (moebius_pt M v) = moebius_circline M (poincare_line u v)" proof (rule unique_poincare_line[symmetric]) show "moebius_pt M u \ moebius_pt M v" using \u \ v\ by auto next show "moebius_pt M u \ circline_set (moebius_circline M (poincare_line u v))" "moebius_pt M v \ circline_set (moebius_circline M (poincare_line u v))" unfolding circline_set_def using moebius_circline[of M "poincare_line u v"] \u \ v\ by auto next from assms(1) have "unit_circle_fix M" by simp thus "is_poincare_line (moebius_circline M (poincare_line u v))" using unit_circle_fix_preserve_is_poincare_line assms by auto next show "moebius_pt M u \ unit_disc" "moebius_pt M v \ unit_disc" using assms(2-3) unit_disc_fix_iff[OF assms(1)] by auto qed text\Conjugate preserve the h-line construction\ lemma conjugate_preserve_poincare_line [simp]: assumes "u \ unit_disc" "v \ unit_disc" "u \ v" shows "poincare_line (conjugate u) (conjugate v) = conjugate_circline (poincare_line u v)" proof- have "conjugate u \ conjugate v" using \u \ v\ by (auto simp add: conjugate_inj) moreover have "conjugate u \ unit_disc" "conjugate v \ unit_disc" using assms by auto moreover have "conjugate u \ circline_set (conjugate_circline (poincare_line u v))" "conjugate v \ circline_set (conjugate_circline (poincare_line u v))" using \u \ v\ by simp_all moreover have "is_poincare_line (conjugate_circline (poincare_line u v))" using is_poincare_line_poincare_line[OF \u \ v\] by simp ultimately show ?thesis using unique_poincare_line[of "conjugate u" "conjugate v" "conjugate_circline (poincare_line u v)"] by simp qed subsubsection \Collinear points and h-lines\ lemma poincare_collinear3_poincare_line_general: assumes "poincare_collinear {a, a1, a2}" "a1 \ a2" "a1 \ inversion a2" shows "a \ circline_set (poincare_line a1 a2)" using assms using poincare_collinear_def unique_poincare_line_general by auto lemma poincare_line_poincare_collinear3_general: assumes "a \ circline_set (poincare_line a1 a2)" "a1 \ a2" shows "poincare_collinear {a, a1, a2}" using assms unfolding poincare_collinear_def by (rule_tac x="poincare_line a1 a2" in exI, simp) lemma poincare_collinear3_poincare_lines_equal_general: assumes "poincare_collinear {a, a1, a2}" "a \ a1" "a \ a2" "a \ inversion a1" "a \ inversion a2" shows "poincare_line a a1 = poincare_line a a2" using assms using unique_poincare_line_general[of a a2 "poincare_line a a1"] by (simp add: insert_commute poincare_collinear3_poincare_line_general) subsubsection \Points collinear with @{term "0\<^sub>h"}\ lemma poincare_collinear_zero_iff: assumes "of_complex y' \ unit_disc" and "of_complex z' \ unit_disc" and "y' \ z'" and "y' \ 0" and "z' \ 0" shows "poincare_collinear {0\<^sub>h, of_complex y', of_complex z'} \ y'*cnj z' = cnj y'*z'" (is "?lhs \ ?rhs") proof- have "of_complex y' \ of_complex z'" using assms using of_complex_inj by blast show ?thesis proof assume ?lhs hence "0\<^sub>h \ circline_set (poincare_line (of_complex y') (of_complex z'))" using unique_poincare_line[of "of_complex y'" "of_complex z'"] using assms \of_complex y' \ of_complex z'\ unfolding poincare_collinear_def by auto moreover let ?mix = "y' * cnj z' - cnj y' * z'" have "is_real (\ * ?mix)" using eq_cnj_iff_real[of ?mix] by auto hence "y' * cnj z' = cnj y' * z' \ Re (\ * ?mix) = 0" using complex.expand[of "\ * ?mix" 0] by (metis complex_i_not_zero eq_iff_diff_eq_0 mult_eq_0_iff zero_complex.simps(1) zero_complex.simps(2)) ultimately show ?rhs using \y' \ z'\ \y' \ 0\ \z' \ 0\ unfolding circline_set_def by simp (transfer, transfer, auto simp add: vec_cnj_def split: if_split_asm, metis Re_complex_of_real Re_mult_real Im_complex_of_real) next assume ?rhs thus ?lhs using assms \of_complex y' \ of_complex z'\ unfolding poincare_collinear_def unfolding circline_set_def apply (rule_tac x="poincare_line (of_complex y') (of_complex z')" in exI) apply auto apply (transfer, transfer, simp add: vec_cnj_def) done qed qed lemma poincare_collinear_zero_polar_form: assumes "poincare_collinear {0\<^sub>h, of_complex x, of_complex y}" and "x \ 0" and "y \ 0" and "of_complex x \ unit_disc" and "of_complex y \ unit_disc" shows "\ \ rx ry. x = cor rx * cis \ \ y = cor ry * cis \ \ rx \ 0 \ ry \ 0" proof- from \x \ 0\ \y \ 0\ obtain \ \' rx ry where - polar: "x = cor rx * cis \" "y = cor ry * cis \'" and "\ = arg x" "\' = arg y" + polar: "x = cor rx * cis \" "y = cor ry * cis \'" and "\ = Arg x" "\' = Arg y" by (metis cmod_cis) hence "rx \ 0" "ry \ 0" using \x \ 0\ \y \ 0\ by auto have "of_complex y \ circline_set (poincare_line 0\<^sub>h (of_complex x))" using assms using unique_poincare_line[of "0\<^sub>h" "of_complex x"] unfolding poincare_collinear_def unfolding circline_set_def using of_complex_zero_iff by fastforce hence "cnj x * y = x * cnj y" using \x \ 0\ \y \ 0\ unfolding circline_set_def by simp (transfer, transfer, simp add: vec_cnj_def field_simps) hence "cis(\' - \) = cis(\ - \')" using polar \rx \ 0\ \ry \ 0\ by (simp add: cis_mult) hence "sin (\' - \) = 0" using cis_diff_cis_opposite[of "\' - \"] by simp then obtain k :: int where "\' - \ = k * pi" using sin_zero_iff_int2[of "\' - \"] by auto hence *: "\' = \ + k * pi" by simp show ?thesis proof (cases "even k") case True then obtain k' where "k = 2*k'" using evenE by blast hence "cis \ = cis \'" using * cos_periodic_int sin_periodic_int by (simp add: cis.ctr field_simps) thus ?thesis using polar \rx \ 0\ \ry \ 0\ by (rule_tac x=\ in exI, rule_tac x=rx in exI, rule_tac x=ry in exI) simp next case False then obtain k' where "k = 2*k' + 1" using oddE by blast hence "cis \ = - cis \'" using * cos_periodic_int sin_periodic_int by (simp add: cis.ctr complex_minus field_simps) thus ?thesis using polar \rx \ 0\ \ry \ 0\ by (rule_tac x=\ in exI, rule_tac x=rx in exI, rule_tac x="-ry" in exI) simp qed qed end diff --git a/thys/Poincare_Disc/Poincare_Lines_Axis_Intersections.thy b/thys/Poincare_Disc/Poincare_Lines_Axis_Intersections.thy --- a/thys/Poincare_Disc/Poincare_Lines_Axis_Intersections.thy +++ b/thys/Poincare_Disc/Poincare_Lines_Axis_Intersections.thy @@ -1,1234 +1,1234 @@ theory Poincare_Lines_Axis_Intersections imports Poincare_Between begin (* ------------------------------------------------------------------ *) section\Intersection of h-lines with the x-axis in the Poincar\'e model\ (* ------------------------------------------------------------------ *) (* ---------------------------------------------------------------- *) subsection\Betweeness of x-axis intersection\ (* ---------------------------------------------------------------- *) text\The intersection point of the h-line determined by points $u$ and $v$ and the x-axis is between $u$ and $v$, then $u$ and $v$ are in the opposite half-planes (one must be in the upper, and the other one in the lower half-plane).\ lemma poincare_between_x_axis_intersection: assumes "u \ unit_disc" and "v \ unit_disc" and "z \ unit_disc" and "u \ v" assumes "u \ circline_set x_axis" and "v \ circline_set x_axis" assumes "z \ circline_set (poincare_line u v) \ circline_set x_axis" - shows "poincare_between u z v \ arg (to_complex u) * arg (to_complex v) < 0" + shows "poincare_between u z v \ Arg (to_complex u) * Arg (to_complex v) < 0" proof- have "\ u v. u \ unit_disc \ v \ unit_disc \ u \ v \ u \ circline_set x_axis \ v \ circline_set x_axis \ z \ circline_set (poincare_line u v) \ circline_set x_axis \ - (poincare_between u z v \ arg (to_complex u) * arg (to_complex v) < 0)" (is "?P z") + (poincare_between u z v \ Arg (to_complex u) * Arg (to_complex v) < 0)" (is "?P z") proof (rule wlog_real_zero) show "?P 0\<^sub>h" proof ((rule allI)+, rule impI, (erule conjE)+) fix u v assume *: "u \ unit_disc" "v \ unit_disc" "u \ v" "u \ circline_set x_axis" "v \ circline_set x_axis" "0\<^sub>h \ circline_set (poincare_line u v) \ circline_set x_axis" obtain u' v' where uv: "u = of_complex u'" "v = of_complex v'" using * inf_or_of_complex[of u] inf_or_of_complex[of v] by auto hence "u \ 0\<^sub>h" "v \ 0\<^sub>h" "u' \ 0" "v' \ 0" using * by auto - hence "arg u' \ 0" "arg v' \ 0" + hence "Arg u' \ 0" "Arg v' \ 0" using * arg_0_iff[of u'] arg_0_iff[of v'] unfolding circline_set_x_axis uv by auto have "poincare_collinear {0\<^sub>h, u, v}" using * unfolding poincare_collinear_def by (rule_tac x="poincare_line u v" in exI, simp) - have "(\k<0. u' = cor k * v') \ (arg u' * arg v' < 0)" (is "?lhs \ ?rhs") + have "(\k<0. u' = cor k * v') \ (Arg u' * Arg v' < 0)" (is "?lhs \ ?rhs") proof assume "?lhs" then obtain k where "k < 0" "u' = cor k * v'" by auto thus ?rhs using arg_mult_real_negative[of k v'] arg_uminus_opposite_sign[of v'] - using \u' \ 0\ \v' \ 0\ \arg u' \ 0\ \arg v' \ 0\ + using \u' \ 0\ \v' \ 0\ \Arg u' \ 0\ \Arg v' \ 0\ by (auto simp add: mult_neg_pos mult_pos_neg) next assume ?rhs obtain ru rv \ where polar: "u' = cor ru * cis \" "v' = cor rv * cis \" using \poincare_collinear {0\<^sub>h, u, v}\ poincare_collinear_zero_polar_form[of u' v'] uv * \u' \ 0\ \v' \ 0\ by auto have "ru * rv < 0" using polar \?rhs\ \u' \ 0\ \v' \ 0\ using arg_mult_real_negative[of "ru" "cis \"] arg_mult_real_positive[of "ru" "cis \"] using arg_mult_real_negative[of "rv" "cis \"] arg_mult_real_positive[of "rv" "cis \"] apply (cases "ru > 0") apply (cases "rv > 0", simp, simp add: mult_pos_neg) apply (cases "rv > 0", simp add: mult_neg_pos, simp) done thus "?lhs" using polar by (rule_tac x="ru / rv" in exI, auto simp add: divide_less_0_iff mult_less_0_iff) qed - thus "poincare_between u 0\<^sub>h v = (arg (to_complex u) * arg (to_complex v) < 0)" + thus "poincare_between u 0\<^sub>h v = (Arg (to_complex u) * Arg (to_complex v) < 0)" using poincare_between_u0v[of u v] * \u \ 0\<^sub>h\ \v \ 0\<^sub>h\ uv by simp qed next fix a z assume 1: "is_real a" "cmod a < 1" "z \ unit_disc" assume 2: "?P (moebius_pt (blaschke a) z)" show "?P z" proof ((rule allI)+, rule impI, (erule conjE)+) fix u v let ?M = "moebius_pt (blaschke a)" let ?Mu = "?M u" let ?Mv = "?M v" assume *: "u \ unit_disc" "v \ unit_disc" "u \ v" "u \ circline_set x_axis" "v \ circline_set x_axis" hence "u \ \\<^sub>h" "v \ \\<^sub>h" by auto have **: "\ x y :: real. x * y < 0 \ sgn (x * y) < 0" by simp assume "z \ circline_set (poincare_line u v) \ circline_set x_axis" - thus "poincare_between u z v = (arg (to_complex u) * arg (to_complex v) < 0)" + thus "poincare_between u z v = (Arg (to_complex u) * Arg (to_complex v) < 0)" using * 1 2[rule_format, of ?Mu ?Mv] \cmod a < 1\ \is_real a\ blaschke_unit_disc_fix[of a] using inversion_noteq_unit_disc[of "of_complex a" u] \u \ \\<^sub>h\ using inversion_noteq_unit_disc[of "of_complex a" v] \v \ \\<^sub>h\ apply auto apply (subst (asm) **, subst **, subst (asm) sgn_mult, subst sgn_mult, simp) apply (subst (asm) **, subst (asm) **, subst (asm) sgn_mult, subst (asm) sgn_mult, simp) done qed next show "z \ unit_disc" by fact next show "is_real (to_complex z)" using assms inf_or_of_complex[of z] by (auto simp add: circline_set_x_axis) qed thus ?thesis using assms by simp qed (* ------------------------------------------------------------------ *) subsection\Check if an h-line intersects the x-axis\ (* ------------------------------------------------------------------ *) lemma x_axis_intersection_equation: assumes "H = mk_circline A B C D" and "(A, B, C, D) \ hermitean_nonzero" shows "of_complex z \ circline_set x_axis \ circline_set H \ A*z\<^sup>2 + 2*Re B*z + D = 0 \ is_real z" (is "?lhs \ ?rhs") proof- have "?lhs \ A*z\<^sup>2 + (B + cnj B)*z + D = 0 \ z = cnj z" using assms using circline_equation_x_axis[of z] using circline_equation[of H A B C D z] using hermitean_elems by (auto simp add: power2_eq_square field_simps) thus ?thesis using eq_cnj_iff_real[of z] using hermitean_elems[of A B C D] by (simp add: complex_add_cnj complex_eq_if_Re_eq) qed text \Check if an h-line intersects x-axis within the unit disc - this could be generalized to checking if an arbitrary circline intersects the x-axis, but we do not need that.\ definition intersects_x_axis_cmat :: "complex_mat \ bool" where [simp]: "intersects_x_axis_cmat H = (let (A, B, C, D) = H in A = 0 \ (Re B)\<^sup>2 > (Re A)\<^sup>2)" lift_definition intersects_x_axis_clmat :: "circline_mat \ bool" is intersects_x_axis_cmat done lift_definition intersects_x_axis :: "circline \ bool" is intersects_x_axis_clmat proof (transfer) fix H1 H2 assume hh: "hermitean H1 \ H1 \ mat_zero" and "hermitean H2 \ H2 \ mat_zero" obtain A1 B1 C1 D1 A2 B2 C2 D2 where *: "H1 = (A1, B1, C1, D1)" "H2 = (A2, B2, C2, D2)" by (cases H1, cases H2, auto) assume "circline_eq_cmat H1 H2" then obtain k where k: "k \ 0 \ H2 = cor k *\<^sub>s\<^sub>m H1" by auto show "intersects_x_axis_cmat H1 = intersects_x_axis_cmat H2" proof- have "k \ 0 \ (Re A1)\<^sup>2 < (Re B1)\<^sup>2 \ (k * Re A1)\<^sup>2 < (k * Re B1)\<^sup>2" by (smt mult_strict_left_mono power2_eq_square semiring_normalization_rules(13) zero_less_power2) thus ?thesis using * k by auto qed qed lemma intersects_x_axis_mk_circline: assumes "is_real A" and "A \ 0 \ B \ 0" shows "intersects_x_axis (mk_circline A B (cnj B) A) \ A = 0 \ (Re B)\<^sup>2 > (Re A)\<^sup>2" proof- let ?H = "(A, B, (cnj B), A)" have "hermitean ?H" using \is_real A\ unfolding hermitean_def mat_adj_def mat_cnj_def using eq_cnj_iff_real by auto moreover have "?H \ mat_zero" using assms by auto ultimately show ?thesis by (transfer, transfer, auto simp add: Let_def) qed lemma intersects_x_axis_iff: assumes "is_poincare_line H" shows "(\ x \ unit_disc. x \ circline_set H \ circline_set x_axis) \ intersects_x_axis H" proof- obtain Ac B C Dc where *: "H = mk_circline Ac B C Dc" "hermitean (Ac, B, C, Dc)" "(Ac, B, C, Dc) \ mat_zero" using ex_mk_circline[of H] by auto hence "(cmod B)\<^sup>2 > (cmod Ac)\<^sup>2" "Ac = Dc" using assms using is_poincare_line_mk_circline by auto hence "H = mk_circline (Re Ac) B (cnj B) (Re Ac)" "hermitean (cor (Re Ac), B, (cnj B), cor (Re Ac))" "(cor (Re Ac), B, (cnj B), cor (Re Ac)) \ mat_zero" using hermitean_elems[of Ac B C Dc] * by auto then obtain A where *: "H = mk_circline (cor A) B (cnj B) (cor A)" "(cor A, B, (cnj B), cor A) \ hermitean_nonzero" by auto show ?thesis proof (cases "A = 0") case True thus ?thesis using * using x_axis_intersection_equation[OF *(1-2), of 0] using intersects_x_axis_mk_circline[of "cor A" B] by auto next case False show ?thesis proof assume "\ x \ unit_disc. x \ circline_set H \ circline_set x_axis" then obtain x where **: "of_complex x \ unit_disc" "of_complex x \ circline_set H \ circline_set x_axis" by (metis inf_or_of_complex inf_notin_unit_disc) hence "is_real x" unfolding circline_set_x_axis using of_complex_inj by auto hence eq: "A * (Re x)\<^sup>2 + 2 * Re B * Re x + A = 0" using ** using x_axis_intersection_equation[OF *(1-2), of "Re x"] by simp hence "(2 * Re B)\<^sup>2 - 4 * A * A \ 0" using discriminant_iff[of A _ "2 * Re B" A] using discrim_def[of A "2 * Re B" A] False by auto hence "(Re B)\<^sup>2 \ (Re A)\<^sup>2" by (simp add: power2_eq_square) moreover have "(Re B)\<^sup>2 \ (Re A)\<^sup>2" proof (rule ccontr) assume "\ ?thesis" hence "Re B = Re A \ Re B = - Re A" using power2_eq_iff by blast hence "A * (Re x)\<^sup>2 + A * 2* Re x + A = 0 \ A * (Re x)\<^sup>2 - A * 2 * Re x + A = 0" using eq by auto hence "A * ((Re x)\<^sup>2 + 2* Re x + 1) = 0 \ A * ((Re x)\<^sup>2 - 2 * Re x + 1) = 0" by (simp add: field_simps) hence "(Re x)\<^sup>2 + 2 * Re x + 1 = 0 \ (Re x)\<^sup>2 - 2 * Re x + 1 = 0" using \A \ 0\ by simp hence "(Re x + 1)\<^sup>2 = 0 \ (Re x - 1)\<^sup>2 = 0" by (simp add: power2_sum power2_diff field_simps) hence "Re x = -1 \ Re x = 1" by auto thus False using \is_real x\ \of_complex x \ unit_disc\ by (auto simp add: cmod_eq_Re) qed ultimately show "intersects_x_axis H" using intersects_x_axis_mk_circline using * by auto next assume "intersects_x_axis H" hence "(Re B)\<^sup>2 > (Re A)\<^sup>2" using * False using intersects_x_axis_mk_circline by simp hence discr: "(2 * Re B)\<^sup>2 - 4 * A * A > 0" by (simp add: power2_eq_square) then obtain x1 x2 where eqs: "A * x1\<^sup>2 + 2 * Re B * x1 + A = 0" "A * x2\<^sup>2 + 2 * Re B * x2 + A = 0" "x1 \ x2" using discriminant_pos_ex[OF \A \ 0\, of "2 * Re B" A] using discrim_def[of A "2 * Re B" A] by auto hence "x1 * x2 = 1" using viette2[OF \A \ 0\, of "2 * Re B" A x1 x2] discr \A \ 0\ by auto have "abs x1 \ 1" "abs x2 \ 1" using eqs discr \x1 * x2 = 1\ by (auto simp add: abs_if power2_eq_square) hence "abs x1 < 1 \ abs x2 < 1" using \x1 * x2 = 1\ by (smt mult_le_cancel_left1 mult_minus_right) thus "\x \ unit_disc. x \ circline_set H \ circline_set x_axis" using x_axis_intersection_equation[OF *(1-2), of x1] using x_axis_intersection_equation[OF *(1-2), of x2] using eqs by auto qed qed qed (* ------------------------------------------------------------------ *) subsection\Check if a Poincar\'e line intersects the y-axis\ (* ------------------------------------------------------------------ *) definition intersects_y_axis_cmat :: "complex_mat \ bool" where [simp]: "intersects_y_axis_cmat H = (let (A, B, C, D) = H in A = 0 \ (Im B)\<^sup>2 > (Re A)\<^sup>2)" lift_definition intersects_y_axis_clmat :: "circline_mat \ bool" is intersects_y_axis_cmat done lift_definition intersects_y_axis :: "circline \ bool" is intersects_y_axis_clmat proof (transfer) fix H1 H2 assume hh: "hermitean H1 \ H1 \ mat_zero" and "hermitean H2 \ H2 \ mat_zero" obtain A1 B1 C1 D1 A2 B2 C2 D2 where *: "H1 = (A1, B1, C1, D1)" "H2 = (A2, B2, C2, D2)" by (cases H1, cases H2, auto) assume "circline_eq_cmat H1 H2" then obtain k where k: "k \ 0 \ H2 = cor k *\<^sub>s\<^sub>m H1" by auto show "intersects_y_axis_cmat H1 = intersects_y_axis_cmat H2" proof- have "k \ 0 \ (Re A1)\<^sup>2 < (Im B1)\<^sup>2 \ (k * Re A1)\<^sup>2 < (k * Im B1)\<^sup>2" by (smt mult_strict_left_mono power2_eq_square semiring_normalization_rules(13) zero_less_power2) thus ?thesis using * k by auto qed qed lemma intersects_x_axis_intersects_y_axis [simp]: shows "intersects_x_axis (moebius_circline (moebius_rotation (pi/2)) H) \ intersects_y_axis H" unfolding moebius_rotation_def moebius_similarity_def by simp (transfer, transfer, auto simp add: mat_adj_def mat_cnj_def) lemma intersects_y_axis_iff: assumes "is_poincare_line H" shows "(\ y \ unit_disc. y \ circline_set H \ circline_set y_axis) \ intersects_y_axis H" (is "?lhs \ ?rhs") proof- let ?R = "moebius_rotation (pi / 2)" let ?H' = "moebius_circline ?R H" have 1: "is_poincare_line ?H'" using assms using unit_circle_fix_preserve_is_poincare_line[OF _ assms, of ?R] by simp show ?thesis proof assume "?lhs" then obtain y where "y \ unit_disc" "y \ circline_set H \ circline_set y_axis" by auto hence "moebius_pt ?R y \ unit_disc \ moebius_pt ?R y \ circline_set ?H' \ circline_set x_axis" using rotation_pi_2_y_axis by (metis Int_iff circline_set_moebius_circline_E moebius_circline_comp_inv_left moebius_pt_comp_inv_left unit_disc_fix_discI unit_disc_fix_rotation) thus ?rhs using intersects_x_axis_iff[OF 1] using intersects_x_axis_intersects_y_axis[of H] by auto next assume "intersects_y_axis H" hence "intersects_x_axis ?H'" using intersects_x_axis_intersects_y_axis[of H] by simp then obtain x where *: "x \ unit_disc" "x \ circline_set ?H' \ circline_set x_axis" using intersects_x_axis_iff[OF 1] by auto let ?y = "moebius_pt (-?R) x" have "?y \ unit_disc \ ?y \ circline_set H \ circline_set y_axis" using * rotation_pi_2_y_axis[symmetric] by (metis Int_iff circline_set_moebius_circline_E moebius_pt_comp_inv_left moebius_rotation_uminus uminus_moebius_def unit_disc_fix_discI unit_disc_fix_rotation) thus ?lhs by auto qed qed (* ------------------------------------------------------------------ *) subsection\Intersection point of a Poincar\'e line with the x-axis in the unit disc\ (* ------------------------------------------------------------------ *) definition calc_x_axis_intersection_cvec :: "complex \ complex \ complex_vec" where [simp]: "calc_x_axis_intersection_cvec A B = (let discr = (Re B)\<^sup>2 - (Re A)\<^sup>2 in (-Re(B) + sgn (Re B) * sqrt(discr), A))" (* intersection with the x-axis for poincare lines that are euclidean circles *) definition calc_x_axis_intersection_cmat_cvec :: "complex_mat \ complex_vec" where [simp]: "calc_x_axis_intersection_cmat_cvec H = (let (A, B, C, D) = H in if A \ 0 then calc_x_axis_intersection_cvec A B else (0, 1) )" lift_definition calc_x_axis_intersection_clmat_hcoords :: "circline_mat \ complex_homo_coords" is calc_x_axis_intersection_cmat_cvec by (auto split: if_split_asm) lift_definition calc_x_axis_intersection :: "circline \ complex_homo" is calc_x_axis_intersection_clmat_hcoords proof transfer fix H1 H2 assume *: "hermitean H1 \ H1 \ mat_zero" "hermitean H2 \ H2 \ mat_zero" 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, auto) assume "circline_eq_cmat H1 H2" then obtain k where k: "k \ 0" "H2 = cor k *\<^sub>s\<^sub>m H1" by auto have "calc_x_axis_intersection_cvec A1 B1 \\<^sub>v calc_x_axis_intersection_cvec A2 B2" using hh k apply simp apply (rule_tac x="cor k" in exI) apply auto apply (simp add: sgn_mult power_mult_distrib) apply (subst right_diff_distrib[symmetric]) apply (subst real_sqrt_mult) by (simp add: real_sgn_eq right_diff_distrib) thus "calc_x_axis_intersection_cmat_cvec H1 \\<^sub>v calc_x_axis_intersection_cmat_cvec H2" using hh k by (auto simp del: calc_x_axis_intersection_cvec_def) qed lemma calc_x_axis_intersection_in_unit_disc: assumes "is_poincare_line H" "intersects_x_axis H" shows "calc_x_axis_intersection H \ unit_disc" proof (cases "is_line H") case True thus ?thesis using assms unfolding unit_disc_def disc_def by simp (transfer, transfer, auto simp add: vec_cnj_def) next case False thus ?thesis using assms unfolding unit_disc_def disc_def proof (simp, transfer, transfer) fix H assume hh: "hermitean H \ H \ mat_zero" then obtain A B D where *: "H = (A, B, cnj B, D)" "is_real A" "is_real D" using hermitean_elems by (cases H) blast assume "is_poincare_line_cmat H" hence *: "H = (A, B, cnj B, A)" "is_real A" using * by auto assume "\ circline_A0_cmat H" hence "A \ 0" using * by simp assume "intersects_x_axis_cmat H" hence "(Re B)\<^sup>2 > (Re A)\<^sup>2" using * \A \ 0\ by (auto simp add: power2_eq_square complex.expand) hence "Re B \ 0" by auto have "Re A \ 0" using \is_real A\ \A \ 0\ by (auto simp add: complex.expand) have "sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2) < sqrt((Re B)\<^sup>2)" using \Re A \ 0\ by (subst real_sqrt_less_iff) auto also have "... = sgn (Re B) * (Re B)" by (smt mult_minus_right nonzero_eq_divide_eq real_sgn_eq real_sqrt_abs) finally have 1: "sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2) < sgn (Re B) * (Re B)" . have 2: "(Re B)\<^sup>2 - (Re A)\<^sup>2 < sgn (Re B) * (Re B) * sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2)" using \(Re B)\<^sup>2 > (Re A)\<^sup>2\ using mult_strict_right_mono[OF 1, of "sqrt ((Re B)\<^sup>2 - (Re A)\<^sup>2)"] by simp have 3: "(Re B)\<^sup>2 - 2*sgn (Re B)*Re B*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2) + (Re B)\<^sup>2 - (Re A)\<^sup>2 < (Re A)\<^sup>2" using mult_strict_left_mono[OF 2, of 2] by (simp add: field_simps) have "(sgn (Re B))\<^sup>2 = 1" using \Re B \ 0\ by (simp add: sgn_if) hence "(-Re B + sgn (Re B) * sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2))\<^sup>2 < (Re A)\<^sup>2" using \(Re B)\<^sup>2 > (Re A)\<^sup>2\ 3 by (simp add: power2_diff field_simps) thus "in_ocircline_cmat_cvec unit_circle_cmat (calc_x_axis_intersection_cmat_cvec H)" using * \(Re B)\<^sup>2 > (Re A)\<^sup>2\ by (auto simp add: vec_cnj_def power2_eq_square split: if_split_asm) qed qed lemma calc_x_axis_intersection: assumes "is_poincare_line H" and "intersects_x_axis H" shows "calc_x_axis_intersection H \ circline_set H \ circline_set x_axis" proof (cases "is_line H") case True thus ?thesis using assms unfolding circline_set_def by simp (transfer, transfer, auto simp add: vec_cnj_def) next case False thus ?thesis using assms unfolding circline_set_def proof (simp, transfer, transfer) fix H assume hh: "hermitean H \ H \ mat_zero" then obtain A B D where *: "H = (A, B, cnj B, D)" "is_real A" "is_real D" using hermitean_elems by (cases H) blast assume "is_poincare_line_cmat H" hence *: "H = (A, B, cnj B, A)" "is_real A" using * by auto assume "\ circline_A0_cmat H" hence "A \ 0" using * by auto assume "intersects_x_axis_cmat H" hence "(Re B)\<^sup>2 > (Re A)\<^sup>2" using * \A \ 0\ by (auto simp add: power2_eq_square complex.expand) hence "Re B \ 0" by auto show "on_circline_cmat_cvec H (calc_x_axis_intersection_cmat_cvec H) \ on_circline_cmat_cvec x_axis_cmat (calc_x_axis_intersection_cmat_cvec H)" (is "?P1 \ ?P2") proof show "on_circline_cmat_cvec H (calc_x_axis_intersection_cmat_cvec H)" proof (cases "circline_A0_cmat H") case True thus ?thesis using * \is_poincare_line_cmat H\ \intersects_x_axis_cmat H\ by (simp add: vec_cnj_def) next case False let ?x = "calc_x_axis_intersection_cvec A B" let ?nom = "fst ?x" and ?den = "snd ?x" have x: "?x = (?nom, ?den)" by simp hence "on_circline_cmat_cvec H (calc_x_axis_intersection_cvec A B)" proof (subst *, subst x, subst on_circline_cmat_cvec_circline_equation) have "(sgn(Re B))\<^sup>2 = 1" using \Re B \ 0\ sgn_pos zero_less_power2 by fastforce have "(sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2))\<^sup>2 = (Re B)\<^sup>2 - (Re A)\<^sup>2" using \(Re B)\<^sup>2 > (Re A)\<^sup>2\ by simp have "(-(Re B) + sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2))\<^sup>2 = (-(Re B))\<^sup>2 + (sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2))\<^sup>2 - 2*(Re B)*sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2)" by (simp add: power2_diff) also have "... = (Re B)*(Re B) + (sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2))\<^sup>2 - 2*(Re B)*sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2)" by (simp add: power2_eq_square) also have "... = (Re B)*(Re B) + (sgn(Re B))\<^sup>2*(sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2))\<^sup>2 - 2*(Re B)*sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2)" by (simp add: power_mult_distrib) also have "... = (Re B)*(Re B) + (Re B)\<^sup>2 - (Re A)\<^sup>2 - 2*(Re B)*sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2)" using \(sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2))\<^sup>2 = (Re B)\<^sup>2 - (Re A)\<^sup>2\ \(sgn(Re B))\<^sup>2 = 1\ by simp finally have "(-(Re B) + sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2))\<^sup>2 = (Re B)*(Re B) + (Re B)\<^sup>2 - (Re A)\<^sup>2 - 2*(Re B)*sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2)" by simp have "is_real ?nom" "is_real ?den" using \is_real A\ by simp+ hence "cnj (?nom) = ?nom" "cnj (?den) = ?den" by (simp add:eq_cnj_iff_real)+ hence "A*?nom*(cnj (?nom)) + B*?den*(cnj (?nom)) + (cnj B)*(cnj (?den))*?nom + A*?den*(cnj (?den)) = A*?nom*?nom + B*?den*?nom + (cnj B)*?den*?nom + A*?den*?den" by auto also have "... = A*?nom*?nom + (B + (cnj B))*?den*?nom + A*?den*?den" by (simp add:field_simps) also have "... = A*?nom*?nom + 2*(Re B)*?den*?nom + A*?den*?den" by (simp add:complex_add_cnj) also have "... = A*?nom\<^sup>2 + 2*(Re B)*?den*?nom + A*?den*?den" by (simp add:power2_eq_square) also have "... = A*(-(Re B) + sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2))\<^sup>2 + 2*(Re B)*A*(-(Re B) + sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2)) + A*A*A" unfolding calc_x_axis_intersection_cvec_def by auto also have "... = A*((Re B)*(Re B) + (Re B)\<^sup>2 - (Re A)\<^sup>2 - 2*(Re B)*sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2)) + 2*(Re B)*A*(-(Re B) + sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2)) + A*A*A" using \(-(Re B) + sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2))\<^sup>2 = (Re B)*(Re B) + (Re B)\<^sup>2 - (Re A)\<^sup>2 - 2*(Re B)*sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2)\ by simp also have "... = A*((Re B)*(Re B) + (Re B)\<^sup>2 - A\<^sup>2 - 2*(Re B)*sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2)) + 2*(Re B)*A*(-(Re B) + sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2)) + A*A*A" using \is_real A\ by simp also have "... = 0" apply (simp add:field_simps) by (simp add: power2_eq_square) finally have "A*?nom*(cnj (?nom)) + B*?den*(cnj (?nom)) + (cnj B)*(cnj (?den))*?nom + A*?den*(cnj (?den)) = 0" by simp thus "circline_equation A B (cnj B) A ?nom ?den" by simp qed thus ?thesis using * \is_poincare_line_cmat H\ \intersects_x_axis_cmat H\ by (simp add: vec_cnj_def) qed next show "on_circline_cmat_cvec x_axis_cmat (calc_x_axis_intersection_cmat_cvec H)" using * \is_poincare_line_cmat H\ \intersects_x_axis_cmat H\ \is_real A\ using eq_cnj_iff_real[of A] by (simp add: vec_cnj_def) qed qed qed lemma unique_calc_x_axis_intersection: assumes "is_poincare_line H" and "H \ x_axis" assumes "x \ unit_disc" and "x \ circline_set H \ circline_set x_axis" shows "x = calc_x_axis_intersection H" proof- have *: "intersects_x_axis H" using assms using intersects_x_axis_iff[OF assms(1)] by auto show "x = calc_x_axis_intersection H" using calc_x_axis_intersection[OF assms(1) *] using calc_x_axis_intersection_in_unit_disc[OF assms(1) *] using assms using unique_is_poincare_line[of x "calc_x_axis_intersection H" H x_axis] by auto qed (* ------------------------------------------------------------------ *) subsection\Check if an h-line intersects the positive part of the x-axis\ (* ------------------------------------------------------------------ *) definition intersects_x_axis_positive_cmat :: "complex_mat \ bool" where [simp]: "intersects_x_axis_positive_cmat H = (let (A, B, C, D) = H in Re A \ 0 \ Re B / Re A < -1)" lift_definition intersects_x_axis_positive_clmat :: "circline_mat \ bool" is intersects_x_axis_positive_cmat done lift_definition intersects_x_axis_positive :: "circline \ bool" is intersects_x_axis_positive_clmat proof (transfer) fix H1 H2 assume hh: "hermitean H1 \ H1 \ mat_zero" and "hermitean H2 \ H2 \ mat_zero" obtain A1 B1 C1 D1 A2 B2 C2 D2 where *: "H1 = (A1, B1, C1, D1)" "H2 = (A2, B2, C2, D2)" by (cases H1, cases H2, auto) assume "circline_eq_cmat H1 H2" then obtain k where "k \ 0 \ H2 = cor k *\<^sub>s\<^sub>m H1" by auto thus "intersects_x_axis_positive_cmat H1 = intersects_x_axis_positive_cmat H2" using * by simp qed lemma intersects_x_axis_positive_mk_circline: assumes "is_real A" and "A \ 0 \ B \ 0" shows "intersects_x_axis_positive (mk_circline A B (cnj B) A) \ Re B / Re A < -1" proof- let ?H = "(A, B, (cnj B), A)" have "hermitean ?H" using \is_real A\ unfolding hermitean_def mat_adj_def mat_cnj_def using eq_cnj_iff_real by auto moreover have "?H \ mat_zero" using assms by auto ultimately show ?thesis by (transfer, transfer, auto simp add: Let_def) qed lemma intersects_x_axis_positive_intersects_x_axis [simp]: assumes "intersects_x_axis_positive H" shows "intersects_x_axis H" proof- have "\ a aa. \ Re a \ 0; Re aa / Re a < - 1; \ (Re a)\<^sup>2 < (Re aa)\<^sup>2 \ \ aa = 0 \ a = 0" by (smt less_divide_eq_1_pos one_less_power pos2 power2_minus power_divide zero_less_power2) thus ?thesis using assms apply transfer apply transfer apply (auto simp add: hermitean_def mat_adj_def mat_cnj_def) done qed lemma add_less_abs_positive_iff: fixes a b :: real assumes "abs b < abs a" shows "a + b > 0 \ a > 0" using assms by auto lemma calc_x_axis_intersection_positive_abs': fixes A B :: real assumes "B\<^sup>2 > A\<^sup>2" and "A \ 0" shows "abs (sgn(B) * sqrt(B\<^sup>2 - A\<^sup>2) / A) < abs(-B/A)" proof- from assms have "B \ 0" by auto have "B\<^sup>2 - A\<^sup>2 < B\<^sup>2" using \A \ 0\ by auto hence "sqrt (B\<^sup>2 - A\<^sup>2) < abs B" using real_sqrt_less_iff[of "B\<^sup>2 - A\<^sup>2" "B\<^sup>2"] by simp thus ?thesis using assms \B \ 0\ by (simp add: abs_mult divide_strict_right_mono) qed lemma calc_intersect_x_axis_positive_lemma: assumes "B\<^sup>2 > A\<^sup>2" and "A \ 0" shows "(-B + sgn B * sqrt(B\<^sup>2 - A\<^sup>2)) / A > 0 \ -B/A > 1" proof- have "(-B + sgn B * sqrt(B\<^sup>2 - A\<^sup>2)) / A = -B / A + (sgn B * sqrt(B\<^sup>2 - A\<^sup>2)) / A" using assms by (simp add: field_simps) moreover have "-B / A + (sgn B * sqrt(B\<^sup>2 - A\<^sup>2)) / A > 0 \ - B / A > 0" using add_less_abs_positive_iff[OF calc_x_axis_intersection_positive_abs'[OF assms]] by simp moreover hence "(B/A)\<^sup>2 > 1" using assms by (simp add: power_divide) hence "B/A > 1 \ B/A < -1" by (smt one_power2 pos2 power2_minus power_0 power_strict_decreasing zero_power2) hence "-B / A > 0 \ -B / A > 1" by auto ultimately show ?thesis using assms by auto qed lemma intersects_x_axis_positive_iff': assumes "is_poincare_line H" shows "intersects_x_axis_positive H \ calc_x_axis_intersection H \ unit_disc \ calc_x_axis_intersection H \ circline_set H \ positive_x_axis" (is "?lhs \ ?rhs") proof let ?x = "calc_x_axis_intersection H" assume ?lhs hence "?x \ circline_set x_axis" "?x \ circline_set H" "?x \ unit_disc" using calc_x_axis_intersection_in_unit_disc[OF assms] calc_x_axis_intersection[OF assms] by auto moreover have "Re (to_complex ?x) > 0" using \?lhs\ assms proof (transfer, transfer) fix H assume hh: "hermitean H \ H \ mat_zero" obtain A B C D where *: "H = (A, B, C, D)" by (cases H, auto) assume "intersects_x_axis_positive_cmat H" hence **: "Re B / Re A < - 1" "Re A \ 0" using * by auto have "(Re B)\<^sup>2 > (Re A)\<^sup>2" using ** by (smt divide_less_eq_1_neg divide_minus_left less_divide_eq_1_pos real_sqrt_abs real_sqrt_less_iff right_inverse_eq) have "is_real A" "A \ 0" using hh hermitean_elems * \Re A \ 0\ complex.expand[of A 0] by auto have "(cmod B)\<^sup>2 > (cmod A)\<^sup>2" using \(Re B)\<^sup>2 > (Re A)\<^sup>2\ \is_real A\ by (smt cmod_power2 power2_less_0 zero_power2) have ***: "0 < (- Re B + sgn (Re B) * sqrt ((Re B)\<^sup>2 - (Re A)\<^sup>2)) / Re A" using calc_intersect_x_axis_positive_lemma[of "Re A" "Re B"] ** \(Re B)\<^sup>2 > (Re A)\<^sup>2\ by auto assume "is_poincare_line_cmat H" hence "A = D" using * hh by simp have "Re ((cor (sgn (Re B)) * cor (sqrt ((Re B)\<^sup>2 - (Re A)\<^sup>2)) - cor (Re B)) / A) = (sgn (Re B) * sqrt ((Re B)\<^sup>2 - (Re A)\<^sup>2) - Re B) / Re D" using \is_real A\ \A = D\ by (metis (no_types, lifting) Re_complex_of_real complex_of_real_Re of_real_diff of_real_divide of_real_mult) thus "0 < Re (to_complex_cvec (calc_x_axis_intersection_cmat_cvec H))" using * hh ** *** \(cmod B)\<^sup>2 > (cmod A)\<^sup>2\ \(Re B)\<^sup>2 > (Re A)\<^sup>2\ \A \ 0\ \A = D\ by simp qed ultimately show ?rhs unfolding positive_x_axis_def by auto next let ?x = "calc_x_axis_intersection H" assume ?rhs hence "Re (to_complex ?x) > 0" "?x \ \\<^sub>h" "?x \ circline_set x_axis" "?x \ unit_disc" "?x \ circline_set H" unfolding positive_x_axis_def by auto hence "intersects_x_axis H" using intersects_x_axis_iff[OF assms] by auto thus ?lhs using \Re (to_complex ?x) > 0\ assms proof (transfer, transfer) fix H assume hh: "hermitean H \ H \ mat_zero" obtain A B C D where *: "H = (A, B, C, D)" by (cases H, auto) assume "0 < Re (to_complex_cvec (calc_x_axis_intersection_cmat_cvec H))" "intersects_x_axis_cmat H" "is_poincare_line_cmat H" hence **: "A \ 0" "0 < Re ((cor (sgn (Re B)) * cor (sqrt ((Re B)\<^sup>2 - (Re A)\<^sup>2)) - cor (Re B)) / A)" "A = D" "is_real A" "(Re B)\<^sup>2 > (Re A)\<^sup>2" using * hh hermitean_elems by (auto split: if_split_asm) have "Re A \ 0" using complex.expand[of A 0] \A \ 0\ \is_real A\ by auto have "Re ((cor (sgn (Re B)) * cor (sqrt ((Re B)\<^sup>2 - (Re D)\<^sup>2)) - cor (Re B)) / D) = (sgn (Re B) * sqrt ((Re B)\<^sup>2 - (Re D)\<^sup>2) - Re B) / Re D" using \is_real A\ \A = D\ by (metis (no_types, lifting) Re_complex_of_real complex_of_real_Re of_real_diff of_real_divide of_real_mult) thus "intersects_x_axis_positive_cmat H" using * ** \Re A \ 0\ using calc_intersect_x_axis_positive_lemma[of "Re A" "Re B"] by simp qed qed lemma intersects_x_axis_positive_iff: assumes "is_poincare_line H" and "H \ x_axis" shows "intersects_x_axis_positive H \ (\ x. x \ unit_disc \ x \ circline_set H \ positive_x_axis)" (is "?lhs \ ?rhs") proof assume ?lhs thus ?rhs using intersects_x_axis_positive_iff'[OF assms(1)] by auto next assume ?rhs then obtain x where "x \ unit_disc" "x \ circline_set H \ positive_x_axis" by auto thus ?lhs using unique_calc_x_axis_intersection[OF assms, of x] using intersects_x_axis_positive_iff'[OF assms(1)] unfolding positive_x_axis_def by auto qed (* ------------------------------------------------------------------ *) subsection\Check if an h-line intersects the positive part of the y-axis\ (* ------------------------------------------------------------------ *) definition intersects_y_axis_positive_cmat :: "complex_mat \ bool" where [simp]: "intersects_y_axis_positive_cmat H = (let (A, B, C, D) = H in Re A \ 0 \ Im B / Re A < -1)" lift_definition intersects_y_axis_positive_clmat :: "circline_mat \ bool" is intersects_y_axis_positive_cmat done lift_definition intersects_y_axis_positive :: "circline \ bool" is intersects_y_axis_positive_clmat proof (transfer) fix H1 H2 assume hh: "hermitean H1 \ H1 \ mat_zero" and "hermitean H2 \ H2 \ mat_zero" obtain A1 B1 C1 D1 A2 B2 C2 D2 where *: "H1 = (A1, B1, C1, D1)" "H2 = (A2, B2, C2, D2)" by (cases H1, cases H2, auto) assume "circline_eq_cmat H1 H2" then obtain k where "k \ 0 \ H2 = cor k *\<^sub>s\<^sub>m H1" by auto thus "intersects_y_axis_positive_cmat H1 = intersects_y_axis_positive_cmat H2" using * by simp qed lemma intersects_x_axis_positive_intersects_y_axis_positive [simp]: shows "intersects_x_axis_positive (moebius_circline (moebius_rotation (-pi/2)) H) \ intersects_y_axis_positive H" using hermitean_elems unfolding moebius_rotation_def moebius_similarity_def by simp (transfer, transfer, auto simp add: mat_adj_def mat_cnj_def) lemma intersects_y_axis_positive_iff: assumes "is_poincare_line H" "H \ y_axis" shows "(\ y \ unit_disc. y \ circline_set H \ positive_y_axis) \ intersects_y_axis_positive H" (is "?lhs \ ?rhs") proof- let ?R = "moebius_rotation (-pi / 2)" let ?H' = "moebius_circline ?R H" have 1: "is_poincare_line ?H'" using assms using unit_circle_fix_preserve_is_poincare_line[OF _ assms(1), of ?R] by simp have 2: "moebius_circline ?R H \ x_axis" proof (rule ccontr) assume "\ ?thesis" hence "H = moebius_circline (moebius_rotation (pi/2)) x_axis" using moebius_circline_comp_inv_left[of ?R H] by auto thus False using \H \ y_axis\ by auto qed show ?thesis proof assume "?lhs" then obtain y where "y \ unit_disc" "y \ circline_set H \ positive_y_axis" by auto hence "moebius_pt ?R y \ unit_disc" "moebius_pt ?R y \ circline_set ?H' \ positive_x_axis" using rotation_minus_pi_2_positive_y_axis by auto thus ?rhs using intersects_x_axis_positive_iff[OF 1 2] using intersects_x_axis_positive_intersects_y_axis_positive[of H] by auto next assume "intersects_y_axis_positive H" hence "intersects_x_axis_positive ?H'" using intersects_x_axis_positive_intersects_y_axis_positive[of H] by simp then obtain x where *: "x \ unit_disc" "x \ circline_set ?H' \ positive_x_axis" using intersects_x_axis_positive_iff[OF 1 2] by auto let ?y = "moebius_pt (-?R) x" have "?y \ unit_disc \ ?y \ circline_set H \ positive_y_axis" using * rotation_minus_pi_2_positive_y_axis[symmetric] by (metis Int_iff circline_set_moebius_circline_E image_eqI moebius_pt_comp_inv_image_left moebius_rotation_uminus uminus_moebius_def unit_disc_fix_discI unit_disc_fix_rotation) thus ?lhs by auto qed qed (* ------------------------------------------------------------------ *) subsection\Position of the intersection point in the unit disc\ (* ------------------------------------------------------------------ *) text\Check if the intersection point of one h-line with the x-axis is located more outward the edge of the disc than the intersection point of another h-line.\ definition outward_cmat :: "complex_mat \ complex_mat \ bool" where [simp]: "outward_cmat H1 H2 = (let (A1, B1, C1, D1) = H1; (A2, B2, C2, D2) = H2 in -Re B1/Re A1 \ -Re B2/Re A2)" lift_definition outward_clmat :: "circline_mat \ circline_mat \ bool" is outward_cmat done lift_definition outward :: "circline \ circline \ bool" is outward_clmat apply transfer apply simp apply (case_tac circline_mat1, case_tac circline_mat2, case_tac circline_mat3, case_tac circline_mat4) apply simp apply (erule_tac exE)+ apply (erule_tac conjE)+ apply simp done lemma outward_mk_circline: assumes "is_real A1" and "is_real A2" and "A1 \ 0 \ B1 \ 0" and "A2 \ 0 \ B2 \ 0" shows "outward (mk_circline A1 B1 (cnj B1) A1) (mk_circline A2 B2 (cnj B2) A2) \ - Re B1 / Re A1 \ - Re B2 / Re A2" proof- let ?H1 = "(A1, B1, (cnj B1), A1)" let ?H2 = "(A2, B2, (cnj B2), A2)" have "hermitean ?H1" "hermitean ?H2" using \is_real A1\ \is_real A2\ unfolding hermitean_def mat_adj_def mat_cnj_def using eq_cnj_iff_real by auto moreover have "?H1 \ mat_zero" "?H2 \ mat_zero" using assms by auto ultimately show ?thesis by (transfer, transfer, auto simp add: Let_def) qed lemma calc_x_axis_intersection_fun_mono: fixes x1 x2 :: real assumes "x1 > 1" and "x2 > x1" shows "x1 - sqrt(x1\<^sup>2 - 1) > x2 - sqrt(x2\<^sup>2 - 1)" using assms proof- have *: "sqrt(x1\<^sup>2 - 1) + sqrt(x2\<^sup>2 - 1) > 0" using assms by (smt one_less_power pos2 real_sqrt_gt_zero) have "sqrt(x1\<^sup>2 - 1) < x1" using real_sqrt_less_iff[of "x1\<^sup>2 - 1" "x1\<^sup>2"] \x1 > 1\ by auto moreover have "sqrt(x2\<^sup>2 - 1) < x2" using real_sqrt_less_iff[of "x2\<^sup>2 - 1" "x2\<^sup>2"] \x1 > 1\ \x2 > x1\ by auto ultimately have "sqrt(x1\<^sup>2 - 1) + sqrt(x2\<^sup>2 - 1) < x1 + x2" by simp hence "(x1 + x2) / (sqrt(x1\<^sup>2 - 1) + sqrt(x2\<^sup>2 - 1)) > 1" using * using less_divide_eq_1_pos[of "sqrt(x1\<^sup>2 - 1) + sqrt(x2\<^sup>2 - 1)" "x1 + x2"] by simp hence "(x2\<^sup>2 - x1\<^sup>2) / (sqrt(x1\<^sup>2 - 1) + sqrt(x2\<^sup>2 - 1)) > x2 - x1" using \x2 > x1\ using mult_less_cancel_left_pos[of "x2 - x1" 1 "(x2 + x1) / (sqrt(x1\<^sup>2 - 1) + sqrt(x2\<^sup>2 - 1))"] by (simp add: power2_eq_square field_simps) moreover have "(x2\<^sup>2 - x1\<^sup>2) = (sqrt(x1\<^sup>2 - 1) + sqrt(x2\<^sup>2 - 1)) * ((sqrt(x2\<^sup>2 - 1) - sqrt(x1\<^sup>2 - 1)))" using \x1 > 1\ \x2 > x1\ by (simp add: field_simps) ultimately have "sqrt(x2\<^sup>2 - 1) - sqrt(x1\<^sup>2 - 1) > x2 - x1" using * by simp thus ?thesis by simp qed lemma calc_x_axis_intersection_mono: fixes a1 b1 a2 b2 :: real assumes "-b1/a1 > 1" and "a1 \ 0" and "-b2/a2 \ -b1/a1" and "a2 \ 0" shows "(-b1 + sgn b1 * sqrt(b1\<^sup>2 - a1\<^sup>2)) / a1 \ (-b2 + sgn b2 * sqrt(b2\<^sup>2 - a2\<^sup>2)) / a2" (is "?lhs \ ?rhs") proof- have "?lhs = -b1/a1 - sqrt((-b1/a1)\<^sup>2 - 1)" proof (cases "b1 > 0") case True hence "a1 < 0" using assms by (smt divide_neg_pos) thus ?thesis using \b1 > 0\ \a1 < 0\ by (simp add: real_sqrt_divide field_simps) next case False hence "b1 < 0" using assms by (cases "b1 = 0") auto hence "a1 > 0" using assms by (smt divide_pos_neg) thus ?thesis using \b1 < 0\ \a1 > 0\ by (simp add: real_sqrt_divide field_simps) qed moreover have "?rhs = -b2/a2 - sqrt((-b2/a2)\<^sup>2 - 1)" proof (cases "b2 > 0") case True hence "a2 < 0" using assms by (smt divide_neg_pos) thus ?thesis using \b2 > 0\ \a2 < 0\ by (simp add: real_sqrt_divide field_simps) next case False hence "b2 < 0" using assms by (cases "b2 = 0") auto hence "a2 > 0" using assms by (smt divide_pos_neg) thus ?thesis using \b2 < 0\ \a2 > 0\ by (simp add: real_sqrt_divide field_simps) qed ultimately show ?thesis using calc_x_axis_intersection_fun_mono[of "-b1/a1" "-b2/a2"] using assms by (cases "-b1/a1=-b2/a2", auto) qed lemma outward: assumes "is_poincare_line H1" and "is_poincare_line H2" assumes "intersects_x_axis_positive H1" and "intersects_x_axis_positive H2" assumes "outward H1 H2" shows "Re (to_complex (calc_x_axis_intersection H1)) \ Re (to_complex (calc_x_axis_intersection H2))" proof- have "intersects_x_axis H1" "intersects_x_axis H2" using assms by auto thus ?thesis using assms proof (transfer, transfer) fix H1 H2 assume hh: "hermitean H1 \ H1 \ mat_zero" "hermitean H2 \ H2 \ mat_zero" obtain A1 B1 C1 D1 A2 B2 C2 D2 where *: "H1 = (A1, B1, C1, D1)" "H2 = (A2, B2, C2, D2)" by (cases H1, cases H2, auto) have "is_real A1" "is_real A2" using hermitean_elems * hh by auto assume 1: "intersects_x_axis_positive_cmat H1" "intersects_x_axis_positive_cmat H2" assume 2: "intersects_x_axis_cmat H1" "intersects_x_axis_cmat H2" assume 3: "is_poincare_line_cmat H1" "is_poincare_line_cmat H2" assume 4: "outward_cmat H1 H2" have "A1 \ 0" "A2 \ 0" using * \is_real A1\ \is_real A2\ 1 complex.expand[of A1 0] complex.expand[of A2 0] by auto hence "(sgn (Re B2) * sqrt ((Re B2)\<^sup>2 - (Re A2)\<^sup>2) - Re B2) / Re A2 \ (sgn (Re B1) * sqrt ((Re B1)\<^sup>2 - (Re A1)\<^sup>2) - Re B1) / Re A1" using calc_x_axis_intersection_mono[of "Re B1" "Re A1" "Re B2" "Re A2"] using 1 4 * by simp moreover have "(sgn (Re B2) * sqrt ((Re B2)\<^sup>2 - (Re A2)\<^sup>2) - Re B2) / Re A2 = Re ((cor (sgn (Re B2)) * cor (sqrt ((Re B2)\<^sup>2 - (Re A2)\<^sup>2)) - cor (Re B2)) / A2)" using \is_real A2\ \A2 \ 0\ by (simp add: Re_divide_real) moreover have "(sgn (Re B1) * sqrt ((Re B1)\<^sup>2 - (Re A1)\<^sup>2) - Re B1) / Re A1 = Re ((cor (sgn (Re B1)) * cor (sqrt ((Re B1)\<^sup>2 - (Re A1)\<^sup>2)) - cor (Re B1)) / A1)" using \is_real A1\ \A1 \ 0\ by (simp add: Re_divide_real) ultimately show "Re (to_complex_cvec (calc_x_axis_intersection_cmat_cvec H2)) \ Re (to_complex_cvec (calc_x_axis_intersection_cmat_cvec H1))" using 2 3 \A1 \ 0\ \A2 \ 0\ * \is_real A1\ \is_real A2\ by (simp del: is_poincare_line_cmat_def intersects_x_axis_cmat_def) qed qed (* ------------------------------------------------------------------ *) subsection\Ideal points and x-axis intersection\ (* ------------------------------------------------------------------ *) lemma ideal_points_intersects_x_axis: assumes "is_poincare_line H" and "ideal_points H = {i1, i2}" and "H \ x_axis" shows "intersects_x_axis H \ Im (to_complex i1) * Im (to_complex i2) < 0" using assms proof- have "i1 \ i2" using assms(1) assms(2) ex_poincare_line_points ideal_points_different(1) by blast have "calc_ideal_points H = {i1, i2}" using assms using ideal_points_unique by auto have "\ i1 \ calc_ideal_points H. \ i2 \ calc_ideal_points H. is_poincare_line H \ H \ x_axis \ i1 \ i2 \ (Im (to_complex i1) * Im (to_complex i2) < 0 \ intersects_x_axis H)" proof (transfer, transfer, (rule ballI)+, rule impI, (erule conjE)+, case_tac H, case_tac i1, case_tac i2) fix i11 i12 i21 i22 A B C D H i1 i2 assume H: "H = (A, B, C, D)" "hermitean H" "H \ mat_zero" assume line: "is_poincare_line_cmat H" assume i1: "i1 = (i11, i12)" "i1 \ calc_ideal_points_cmat_cvec H" assume i2: "i2 = (i21, i22)" "i2 \ calc_ideal_points_cmat_cvec H" assume different: "\ i1 \\<^sub>v i2" assume not_x_axis: "\ circline_eq_cmat H x_axis_cmat" have "is_real A" "is_real D" "C = cnj B" using H hermitean_elems by auto have "(cmod A)\<^sup>2 < (cmod B)\<^sup>2" "A = D" using line H by auto let ?discr = "sqrt ((cmod B)\<^sup>2 - (Re D)\<^sup>2)" let ?den = "(cmod B)\<^sup>2" let ?i1 = "B * (- D - \ * ?discr)" let ?i2 = "B * (- D + \ * ?discr)" have "i11 = ?i1 \ i11 = ?i2" "i12 = ?den" "i21 = ?i1 \ i21 = ?i2" "i22 = ?den" using i1 i2 H line by (auto split: if_split_asm) hence i: "i11 = ?i1 \ i21 = ?i2 \ i11 = ?i2 \ i21 = ?i1" using \\ i1 \\<^sub>v i2\ i1 i2 by auto have "Im (i11 / i12) * Im (i21 / i22) = Im (?i1 / ?den) * Im (?i2 / ?den)" using i \i12 = ?den\ \i22 = ?den\ by auto also have "... = Im (?i1) * Im (?i2) / ?den\<^sup>2" by simp also have "... = (Im B * (Im B * (Re D * Re D)) - Re B * (Re B * ((cmod B)\<^sup>2 - (Re D)\<^sup>2))) / cmod B ^ 4" using \(cmod B)\<^sup>2 > (cmod A)\<^sup>2\ \A = D\ using \is_real D\ cmod_eq_Re[of A] by (auto simp add: field_simps) also have "... = ((Im B)\<^sup>2 * (Re D)\<^sup>2 - (Re B)\<^sup>2 * ((Re B)\<^sup>2 + (Im B)\<^sup>2 - (Re D)\<^sup>2)) / cmod B ^ 4" proof- have "cmod B * cmod B = Re B * Re B + Im B * Im B" by (metis cmod_power2 power2_eq_square) thus ?thesis by (simp add: power2_eq_square) qed also have "... = (((Re D)\<^sup>2 - (Re B)\<^sup>2) * ((Re B)\<^sup>2 + (Im B)\<^sup>2)) / cmod B ^ 4" by (simp add: power2_eq_square field_simps) finally have Im_product: "Im (i11 / i12) * Im (i21 / i22) = ((Re D)\<^sup>2 - (Re B)\<^sup>2) * ((Re B)\<^sup>2 + (Im B)\<^sup>2) / cmod B ^ 4" . show "Im (to_complex_cvec i1) * Im (to_complex_cvec i2) < 0 \ intersects_x_axis_cmat H" proof safe assume opposite: "Im (to_complex_cvec i1) * Im (to_complex_cvec i2) < 0" show "intersects_x_axis_cmat H" proof- have "((Re D)\<^sup>2 - (Re B)\<^sup>2) * ((Re B)\<^sup>2 + (Im B)\<^sup>2) / cmod B ^ 4 < 0" using Im_product opposite i1 i2 by simp hence "((Re D)\<^sup>2 - (Re B)\<^sup>2) * ((Re B)\<^sup>2 + (Im B)\<^sup>2) < 0" by (simp add: divide_less_0_iff) hence "(Re D)\<^sup>2 < (Re B)\<^sup>2" by (simp add: mult_less_0_iff not_sum_power2_lt_zero) thus ?thesis using H \A = D\ \is_real D\ by auto qed next have *: "(\k. k * Im B = 1 \ k = 0) \ Im B = 0" apply (safe, erule_tac x="1 / Im B" in allE) using divide_cancel_left by fastforce assume "intersects_x_axis_cmat H" hence "Re D = 0 \ (Re D)\<^sup>2 < (Re B)\<^sup>2" using H \A = D\ by auto hence "(Re D)\<^sup>2 < (Re B)\<^sup>2" using \is_real D\ line H \C = cnj B\ using not_x_axis * by (auto simp add: complex_eq_iff) hence "((Re D)\<^sup>2 - (Re B)\<^sup>2) * ((Re B)\<^sup>2 + (Im B)\<^sup>2) < 0" by (metis add_cancel_left_left diff_less_eq mult_eq_0_iff mult_less_0_iff power2_eq_square power2_less_0 sum_squares_gt_zero_iff) thus "Im (to_complex_cvec i1) * Im (to_complex_cvec i2) < 0" using Im_product i1 i2 using divide_eq_0_iff divide_less_0_iff prod.simps(2) to_complex_cvec_def zero_complex.simps(1) zero_less_norm_iff by fastforce qed qed thus ?thesis using assms \calc_ideal_points H = {i1, i2}\ \i1 \ i2\ by auto qed end 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'" + 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'" + "\ = 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'" + 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'\ * + 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 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)+ have theta_sin: "\ > 0 \ sin \ > 0" "\ < 0 \ sin \ < 0" - using \\ = arg y'\ \\ \ 0\ \\ \ pi\ + using \\ = Arg y'\ \\ \ 0\ \\ \ pi\ 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) + 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\ + 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)" + 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'" + 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) 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) 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" + 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