Theory Hyperbolic_Tarski

(*  Title:       The hyperbolic plane and Tarski's axioms
    Author:      Tim Makarios <tjm1983 at gmail.com>, 2012
    Maintainer:  Tim Makarios <tjm1983 at gmail.com>
*)

section "The hyperbolic plane and Tarski's axioms"

theory Hyperbolic_Tarski
imports Euclid_Tarski
  Projective
  "HOL-Library.Quadratic_Discriminant"
begin

subsection ‹Characterizing a specific conic in the projective plane›

definition M :: "real^3^3" where
  "M  vector [
  vector [1, 0, 0],
  vector [0, 1, 0],
  vector [0, 0, -1]]"

lemma M_symmatrix: "symmatrix M"
  unfolding symmatrix_def and transpose_def and M_def
  by (simp add: vec_eq_iff forall_3 vector_3)

lemma M_self_inverse: "M ** M = mat 1"
  unfolding M_def and matrix_matrix_mult_def and mat_def and vector_def
  by (simp add: sum_3 vec_eq_iff forall_3)

lemma M_invertible: "invertible M"
  unfolding invertible_def
  using M_self_inverse
  by auto

definition polar :: "proj2  proj2_line" where
  "polar p  proj2_line_abs (M *v proj2_rep p)"

definition pole :: "proj2_line  proj2" where
  "pole l  proj2_abs (M *v proj2_line_rep l)"

lemma polar_abs:
  assumes "v  0"
  shows "polar (proj2_abs v) = proj2_line_abs (M *v v)"
proof -
  from v  0 and proj2_rep_abs2
  obtain k where "k  0" and "proj2_rep (proj2_abs v) = k *R v" by auto
  from proj2_rep (proj2_abs v) = k *R v
  have "polar (proj2_abs v) = proj2_line_abs (k *R (M *v v))"
    unfolding polar_def
    by (simp add: matrix_scaleR_vector_ac scaleR_matrix_vector_assoc)
  with k  0 and proj2_line_abs_mult
  show "polar (proj2_abs v) = proj2_line_abs (M *v v)" by simp
qed

lemma pole_abs:
  assumes "v  0"
  shows "pole (proj2_line_abs v) = proj2_abs (M *v v)"
proof -
  from v  0 and proj2_line_rep_abs
  obtain k where "k  0" and "proj2_line_rep (proj2_line_abs v) = k *R v"
    by auto
  from proj2_line_rep (proj2_line_abs v) = k *R v
  have "pole (proj2_line_abs v) = proj2_abs (k *R (M *v v))"
    unfolding pole_def
    by (simp add: matrix_scaleR_vector_ac scaleR_matrix_vector_assoc)
  with k  0 and proj2_abs_mult
  show "pole (proj2_line_abs v) = proj2_abs (M *v v)" by simp
qed

lemma polar_rep_non_zero: "M *v proj2_rep p  0"
proof -
  have "proj2_rep p  0" by (rule proj2_rep_non_zero)
  with M_invertible
  show "M *v proj2_rep p  0" by (rule invertible_times_non_zero)
qed

lemma pole_polar: "pole (polar p) = p"
proof -
  from polar_rep_non_zero
  have "pole (polar p) = proj2_abs (M *v (M *v proj2_rep p))"
    unfolding polar_def
    by (rule pole_abs)
  with M_self_inverse
  show "pole (polar p) = p"
    by (simp add: matrix_vector_mul_assoc proj2_abs_rep)
qed

lemma pole_rep_non_zero: "M *v proj2_line_rep l  0"
proof -
  have "proj2_line_rep l  0" by (rule proj2_line_rep_non_zero)
  with M_invertible
  show "M *v proj2_line_rep l  0" by (rule invertible_times_non_zero)
qed

lemma polar_pole: "polar (pole l) = l"
proof -
  from pole_rep_non_zero
  have "polar (pole l) = proj2_line_abs (M *v (M *v proj2_line_rep l))"
    unfolding pole_def
    by (rule polar_abs)
  with M_self_inverse
  show "polar (pole l) = l"
    by (simp add: matrix_vector_mul_assoc proj2_line_abs_rep
      matrix_vector_mul_lid)
qed

lemma polar_inj:
  assumes "polar p = polar q"
  shows "p = q"
proof -
  from polar p = polar q have "pole (polar p) = pole (polar q)" by simp
  thus "p = q" by (simp add: pole_polar)
qed

definition conic_sgn :: "proj2  real" where
  "conic_sgn p  sgn (proj2_rep p  (M *v proj2_rep p))"

lemma conic_sgn_abs:
  assumes "v  0"
  shows "conic_sgn (proj2_abs v) = sgn (v  (M *v v))"
proof -
  from v  0 and proj2_rep_abs2
  obtain j where "j  0" and "proj2_rep (proj2_abs v) = j *R v" by auto

  from proj2_rep (proj2_abs v) = j *R v
  have "conic_sgn (proj2_abs v) = sgn (j2 * (v  (M *v v)))"
    unfolding conic_sgn_def
    by (simp add:
      matrix_scaleR_vector_ac
      scaleR_matrix_vector_assoc [symmetric]
      dot_scaleR_mult
      power2_eq_square
      algebra_simps)
  also have " = sgn (j2) * sgn (v  (M *v v))" by (rule sgn_mult)
  also from j  0 have " = sgn (v  (M *v v))"
    by (simp add: power2_eq_square sgn_mult)
  finally show "conic_sgn (proj2_abs v) = sgn (v  (M *v v))" .
qed

lemma sgn_conic_sgn: "sgn (conic_sgn p) = conic_sgn p"
  by (unfold conic_sgn_def) simp

definition S :: "proj2 set" where
  "S  {p. conic_sgn p = 0}"

definition K2 :: "proj2 set" where
  "K2  {p. conic_sgn p < 0}"

lemma S_K2_empty: "S  K2 = {}"
  unfolding S_def and K2_def
  by auto

lemma K2_abs:
  assumes "v  0"
  shows "proj2_abs v  K2  v  (M *v v) < 0"
proof -
  have "proj2_abs v  K2  conic_sgn (proj2_abs v) < 0"
    by (simp add: K2_def)
  with v  0 and conic_sgn_abs
  show "proj2_abs v  K2  v  (M *v v) < 0" by simp
qed

definition "K2_centre = proj2_abs (vector [0,0,1])"

lemma K2_centre_non_zero: "vector [0,0,1]  (0 :: real^3)"
  by (unfold vector_def) (simp add: vec_eq_iff forall_3)

lemma K2_centre_in_K2: "K2_centre  K2"
proof -
  from K2_centre_non_zero and proj2_rep_abs2
  obtain k where "k  0" and "proj2_rep K2_centre = k *R vector [0,0,1]"
    by (unfold K2_centre_def) auto
  from k  0 have "0 < k2" by simp
  with proj2_rep K2_centre = k *R vector [0,0,1]
  show "K2_centre  K2"
    unfolding K2_def
      and conic_sgn_def
      and M_def
      and matrix_vector_mult_def
      and inner_vec_def
      and vector_def
    by (simp add: vec_eq_iff sum_3 power2_eq_square)
qed

lemma K2_imp_M_neg:
  assumes "v  0" and "proj2_abs v  K2"
  shows "v  (M *v v) < 0"
  using assms
  by (simp add: K2_abs)

lemma M_neg_imp_z_squared_big:
  assumes "v  (M *v v) < 0"
  shows "(v$3)2 > (v$1)2 + (v$2)2"
  using v  (M *v v) < 0
  unfolding matrix_vector_mult_def and M_def and vector_def
  by (simp add: inner_vec_def sum_3 power2_eq_square)

lemma M_neg_imp_z_non_zero:
  assumes "v  (M *v v) < 0"
  shows "v$3  0"
proof -
  have "(v$1)2 + (v$2)2  0" by simp
  with M_neg_imp_z_squared_big [of v] and v  (M *v v) < 0
  have "(v$3)2 > 0" by arith
  thus "v$3  0" by simp
qed

lemma M_neg_imp_K2:
  assumes "v  (M *v v) < 0"
  shows "proj2_abs v  K2"
proof -
  from v  (M *v v) < 0 have "v$3  0" by (rule M_neg_imp_z_non_zero)
  hence "v  0" by auto
  with v  (M *v v) < 0 and K2_abs show "proj2_abs v  K2" by simp
qed

lemma M_reverse: "a  (M *v b) = b  (M *v a)"
  unfolding matrix_vector_mult_def and M_def and vector_def
  by (simp add: inner_vec_def sum_3)

lemma S_abs:
  assumes "v  0"
  shows "proj2_abs v  S  v  (M *v v) = 0"
proof -
  have "proj2_abs v  S  conic_sgn (proj2_abs v) = 0"
    unfolding S_def
    by simp
  also from v  0 and conic_sgn_abs
  have "  sgn (v  (M *v v)) = 0" by simp
  finally show "proj2_abs v  S  v  (M *v v) = 0" by (simp add: sgn_0_0)
qed

lemma S_alt_def: "p  S  proj2_rep p  (M *v proj2_rep p) = 0"
proof -
  have "proj2_rep p  0" by (rule proj2_rep_non_zero)
  hence "proj2_abs (proj2_rep p)  S  proj2_rep p  (M *v proj2_rep p) = 0"
    by (rule S_abs)
  thus "p  S  proj2_rep p  (M *v proj2_rep p) = 0"
    by (simp add: proj2_abs_rep)
qed

lemma incident_polar:
  "proj2_incident p (polar q)  proj2_rep p  (M *v proj2_rep q) = 0"
  using polar_rep_non_zero
  unfolding polar_def
  by (rule proj2_incident_right_abs)

lemma incident_own_polar_in_S: "proj2_incident p (polar p)  p  S"
  using incident_polar and S_alt_def
  by simp

lemma incident_polar_swap:
  assumes "proj2_incident p (polar q)"
  shows "proj2_incident q (polar p)"
proof -
  from proj2_incident p (polar q)
  have "proj2_rep p  (M *v proj2_rep q) = 0" by (unfold incident_polar)
  hence "proj2_rep q  (M *v proj2_rep p) = 0" by (simp add: M_reverse)
  thus "proj2_incident q (polar p)" by (unfold incident_polar)
qed

lemma incident_pole_polar:
  assumes "proj2_incident p l"
  shows "proj2_incident (pole l) (polar p)"
proof -
  from proj2_incident p l
  have "proj2_incident p (polar (pole l))" by (subst polar_pole)
  thus "proj2_incident (pole l) (polar p)" by (rule incident_polar_swap)
qed

definition z_zero :: "proj2_line" where
  "z_zero  proj2_line_abs (vector [0,0,1])"

lemma z_zero:
  assumes "(proj2_rep p)$3 = 0"
  shows "proj2_incident p z_zero"
proof -
  from K2_centre_non_zero and proj2_line_rep_abs
  obtain k where "proj2_line_rep z_zero = k *R vector [0,0,1]"
    by (unfold z_zero_def) auto
  with (proj2_rep p)$3 = 0
  show "proj2_incident p z_zero"
    unfolding proj2_incident_def and inner_vec_def and vector_def
    by (simp add: sum_3)
qed

lemma z_zero_conic_sgn_1:
  assumes "proj2_incident p z_zero"
  shows "conic_sgn p = 1"
proof -
  let ?v = "proj2_rep p"
  have "(vector [0,0,1] :: real^3)  0"
    unfolding vector_def
    by (simp add: vec_eq_iff)
  with proj2_incident p z_zero
  have "?v  vector [0,0,1] = 0"
    unfolding z_zero_def
    by (simp add: proj2_incident_right_abs)
  hence "?v$3 = 0"
    unfolding inner_vec_def and vector_def
    by (simp add: sum_3)
  hence "?v  (M *v ?v) = (?v$1)2 + (?v$2)2"
    unfolding inner_vec_def
      and power2_eq_square
      and matrix_vector_mult_def
      and M_def
      and vector_def
      and sum_3
    by simp

  have "?v  0" by (rule proj2_rep_non_zero)
  with ?v$3 = 0 have "?v$1  0  ?v$2  0" by (simp add: vec_eq_iff forall_3)
  hence "(?v$1)2 > 0  (?v$2)2 > 0" by simp
  with add_sign_intros [of "(?v$1)2" "(?v$2)2"]
  have "(?v$1)2 + (?v$2)2 > 0" by auto
  with ?v  (M *v ?v) = (?v$1)2 + (?v$2)2
  have "?v  (M *v ?v) > 0" by simp
  thus "conic_sgn p = 1"
    unfolding conic_sgn_def
    by simp
qed

lemma conic_sgn_not_1_z_non_zero:
  assumes "conic_sgn p  1"
  shows "z_non_zero p"
proof -
  from conic_sgn p  1
  have "¬ proj2_incident p z_zero" by (auto simp add: z_zero_conic_sgn_1)
  thus "z_non_zero p" by (auto simp add: z_zero)
qed

lemma z_zero_not_in_S:
  assumes "proj2_incident p z_zero"
  shows "p  S"
proof -
  from proj2_incident p z_zero have "conic_sgn p = 1"
    by (rule z_zero_conic_sgn_1)
  thus "p  S"
    unfolding S_def
    by simp
qed

lemma line_incident_point_not_in_S: " p. p  S  proj2_incident p l"
proof -
  let ?p = "proj2_intersection l z_zero"
  have "proj2_incident ?p l" and "proj2_incident ?p z_zero"
    by (rule proj2_intersection_incident)+
  from proj2_incident ?p z_zero have "?p  S" by (rule z_zero_not_in_S)
  with proj2_incident ?p l
  show " p. p  S  proj2_incident p l" by auto
qed

lemma apply_cltn2_abs_abs_in_S:
  assumes "v  0" and "invertible J"
  shows "apply_cltn2 (proj2_abs v) (cltn2_abs J)  S
   v  (J ** M ** transpose J *v v) = 0"
proof -
  from v  0 and invertible J
  have "v v* J  0" by (rule non_zero_mult_invertible_non_zero)

  from v  0 and invertible J
  have "apply_cltn2 (proj2_abs v) (cltn2_abs J) = proj2_abs (v v* J)"
    by (rule apply_cltn2_abs)
  also from v v* J  0
  have "  S  (v v* J)  (M *v (v v* J)) = 0" by (rule S_abs)
  finally show "apply_cltn2 (proj2_abs v) (cltn2_abs J)  S
     v  (J ** M ** transpose J *v v) = 0"
    by (simp add: dot_lmul_matrix matrix_vector_mul_assoc [symmetric])
qed

lemma apply_cltn2_right_abs_in_S:
  assumes "invertible J"
  shows "apply_cltn2 p (cltn2_abs J)  S
   (proj2_rep p)  (J ** M ** transpose J *v (proj2_rep p)) = 0"
proof -
  have "proj2_rep p  0" by (rule proj2_rep_non_zero)
  with invertible J
  have "apply_cltn2 (proj2_abs (proj2_rep p)) (cltn2_abs J)  S
     proj2_rep p  (J ** M ** transpose J *v proj2_rep p) = 0"
    by (simp add: apply_cltn2_abs_abs_in_S)
  thus "apply_cltn2 p (cltn2_abs J)  S
     proj2_rep p  (J ** M ** transpose J *v proj2_rep p) = 0"
    by (simp add: proj2_abs_rep)
qed

lemma apply_cltn2_abs_in_S:
  assumes "v  0"
  shows "apply_cltn2 (proj2_abs v) C  S
   v  (cltn2_rep C ** M ** transpose (cltn2_rep C) *v v) = 0"
proof -
  have "invertible (cltn2_rep C)" by (rule cltn2_rep_invertible)
  with v  0
  have "apply_cltn2 (proj2_abs v) (cltn2_abs (cltn2_rep C))  S
     v  (cltn2_rep C ** M ** transpose (cltn2_rep C) *v v) = 0"
    by (rule apply_cltn2_abs_abs_in_S)
  thus "apply_cltn2 (proj2_abs v) C  S
     v  (cltn2_rep C ** M ** transpose (cltn2_rep C) *v v) = 0"
    by (simp add: cltn2_abs_rep)
qed

lemma apply_cltn2_in_S:
  "apply_cltn2 p C  S
   proj2_rep p  (cltn2_rep C ** M ** transpose (cltn2_rep C) *v proj2_rep p)
  = 0"
proof -
  have "proj2_rep p  0" by (rule proj2_rep_non_zero)
  hence "apply_cltn2 (proj2_abs (proj2_rep p)) C  S
     proj2_rep p  (cltn2_rep C ** M ** transpose (cltn2_rep C) *v proj2_rep p)
    = 0"
    by (rule apply_cltn2_abs_in_S)
  thus "apply_cltn2 p C  S
     proj2_rep p  (cltn2_rep C ** M ** transpose (cltn2_rep C) *v proj2_rep p)
    = 0"
    by (simp add: proj2_abs_rep)
qed

lemma norm_M: "(vector2_append1 v)  (M *v vector2_append1 v) = (norm v)2 - 1"
proof -
  have "(norm v)2 = (v$1)2 + (v$2)2"
    unfolding norm_vec_def
      and L2_set_def
    by (simp add: sum_2)
  thus "(vector2_append1 v)  (M *v vector2_append1 v) = (norm v)2 - 1"
    unfolding vector2_append1_def
      and inner_vec_def
      and matrix_vector_mult_def
      and vector_def
      and M_def
      and power2_norm_eq_inner
    by (simp add: sum_3 power2_eq_square)
qed

subsection ‹Some specific points and lines of the projective plane›

definition "east = proj2_abs (vector [1,0,1])"
definition "west = proj2_abs (vector [-1,0,1])"
definition "north = proj2_abs (vector [0,1,1])"
definition "south = proj2_abs (vector [0,-1,1])"
definition "far_north = proj2_abs (vector [0,1,0])"

lemmas compass_defs = east_def west_def north_def south_def

lemma compass_non_zero:
  shows "vector [1,0,1]  (0 :: real^3)"
  and "vector [-1,0,1]  (0 :: real^3)"
  and "vector [0,1,1]  (0 :: real^3)"
  and "vector [0,-1,1]  (0 :: real^3)"
  and "vector [0,1,0]  (0 :: real^3)"
  and "vector [1,0,0]  (0 :: real^3)"
  unfolding vector_def
  by (simp_all add: vec_eq_iff forall_3)

lemma east_west_distinct: "east  west"
proof
  assume "east = west"
  with compass_non_zero
    and proj2_abs_abs_mult [of "vector [1,0,1]" "vector [-1,0,1]"]
  obtain k where "(vector [1,0,1] :: real^3) = k *R vector [-1,0,1]"
    unfolding compass_defs
    by auto
  thus False
    unfolding vector_def
    by (auto simp add: vec_eq_iff forall_3)
qed

lemma north_south_distinct: "north  south"
proof
  assume "north = south"
  with compass_non_zero
    and proj2_abs_abs_mult [of "vector [0,1,1]" "vector [0,-1,1]"]
  obtain k where "(vector [0,1,1] :: real^3) = k *R vector [0,-1,1]"
    unfolding compass_defs
    by auto
  thus False
    unfolding vector_def
    by (auto simp add: vec_eq_iff forall_3)
qed

lemma north_not_east_or_west: "north  {east, west}"
proof
  assume "north  {east, west}"
  hence "east = north  west = north" by auto
  with compass_non_zero
    and proj2_abs_abs_mult [of _ "vector [0,1,1]"]
  obtain k where "(vector [1,0,1] :: real^3) = k *R vector [0,1,1]
     (vector [-1,0,1] :: real^3) = k *R vector [0,1,1]"
    unfolding compass_defs
    by auto
  thus False
    unfolding vector_def
    by (simp add: vec_eq_iff forall_3)
qed

lemma compass_in_S:
  shows "east  S" and "west  S" and "north  S" and "south  S"
  using compass_non_zero and S_abs
  unfolding compass_defs
    and M_def
    and inner_vec_def
    and matrix_vector_mult_def
    and vector_def
  by (simp_all add: sum_3)

lemma east_west_tangents:
  shows "polar east = proj2_line_abs (vector [-1,0,1])"
  and "polar west = proj2_line_abs (vector [1,0,1])"
proof -
  have "M *v vector [1,0,1] = (-1) *R vector [-1,0,1]"
    and "M *v vector [-1,0,1] = (-1) *R vector [1,0,1]"
    unfolding M_def and matrix_vector_mult_def and vector_def
    by (simp_all add: vec_eq_iff sum_3)
  with compass_non_zero and polar_abs
  have "polar east = proj2_line_abs ((-1) *R vector [-1,0,1])"
    and "polar west = proj2_line_abs ((-1) *R vector [1,0,1])"
    unfolding compass_defs
    by simp_all
  with proj2_line_abs_mult [of "-1"]
  show "polar east = proj2_line_abs (vector [-1,0,1])"
    and "polar west = proj2_line_abs (vector [1,0,1])"
    by simp_all
qed

lemma east_west_tangents_distinct: "polar east  polar west"
proof
  assume "polar east = polar west"
  hence "east = west" by (rule polar_inj)
  with east_west_distinct show False ..
qed

lemma east_west_tangents_incident_far_north:
  shows "proj2_incident far_north (polar east)"
  and "proj2_incident far_north (polar west)"
  using compass_non_zero and proj2_incident_abs
  unfolding far_north_def and east_west_tangents and inner_vec_def
  by (simp_all add: sum_3 vector_3)

lemma east_west_tangents_far_north:
  "proj2_intersection (polar east) (polar west) = far_north"
  using east_west_tangents_distinct and east_west_tangents_incident_far_north
  by (rule proj2_intersection_unique [symmetric])

instantiation proj2 :: zero
begin
definition proj2_zero_def: "0 = proj2_pt 0"
instance ..
end

definition "equator  proj2_line_abs (vector [0,1,0])"
definition "meridian  proj2_line_abs (vector [1,0,0])"

lemma equator_meridian_distinct: "equator  meridian"
proof
  assume "equator = meridian"
  with compass_non_zero
    and proj2_line_abs_abs_mult [of "vector [0,1,0]" "vector [1,0,0]"]
  obtain k where "(vector [0,1,0] :: real^3) = k *R vector [1,0,0]"
    by (unfold equator_def meridian_def) auto
  thus False by (unfold vector_def) (auto simp add: vec_eq_iff forall_3)
qed

lemma east_west_on_equator:
  shows "proj2_incident east equator" and "proj2_incident west equator"
  unfolding east_def and west_def and equator_def
  using compass_non_zero
  by (simp_all add: proj2_incident_abs inner_vec_def vector_def sum_3)

lemma north_far_north_distinct: "north  far_north"
proof
  assume "north = far_north"
  with compass_non_zero
    and proj2_abs_abs_mult [of "vector [0,1,1]" "vector [0,1,0]"]
  obtain k where "(vector [0,1,1] :: real^3) = k *R vector [0,1,0]"
    by (unfold north_def far_north_def) auto
  thus False
    unfolding vector_def
    by (auto simp add: vec_eq_iff forall_3)
qed

lemma north_south_far_north_on_meridian:
  shows "proj2_incident north meridian" and "proj2_incident south meridian"
  and "proj2_incident far_north meridian"
  unfolding compass_defs and far_north_def and meridian_def
  using compass_non_zero
  by (simp_all add: proj2_incident_abs inner_vec_def vector_def sum_3)

lemma K2_centre_on_equator_meridian:
  shows "proj2_incident K2_centre equator"
  and "proj2_incident K2_centre meridian"
  unfolding K2_centre_def and equator_def and meridian_def
  using K2_centre_non_zero and compass_non_zero
  by (simp_all add: proj2_incident_abs inner_vec_def vector_def sum_3)

lemma on_equator_meridian_is_K2_centre:
  assumes "proj2_incident a equator" and "proj2_incident a meridian"
  shows "a = K2_centre"
  using assms and K2_centre_on_equator_meridian and equator_meridian_distinct
    and proj2_incident_unique
  by auto

definition "rep_equator_reflect  vector [
  vector [1, 0,0],
  vector [0,-1,0],
  vector [0, 0,1]] :: real^3^3"
definition "rep_meridian_reflect  vector [
  vector [-1,0,0],
  vector [ 0,1,0],
  vector [ 0,0,1]] :: real^3^3"
definition "equator_reflect  cltn2_abs rep_equator_reflect"
definition "meridian_reflect  cltn2_abs rep_meridian_reflect"

lemmas compass_reflect_defs = equator_reflect_def meridian_reflect_def
  rep_equator_reflect_def rep_meridian_reflect_def

lemma compass_reflect_self_inverse:
  shows "rep_equator_reflect ** rep_equator_reflect = mat 1"
  and "rep_meridian_reflect ** rep_meridian_reflect = mat 1"
  unfolding compass_reflect_defs matrix_matrix_mult_def mat_def
  by (simp_all add: vec_eq_iff forall_3 sum_3 vector_3)

lemma compass_reflect_invertible:
  shows "invertible rep_equator_reflect" and "invertible rep_meridian_reflect"
  unfolding invertible_def
  using compass_reflect_self_inverse
  by auto

lemma compass_reflect_compass:
  shows "apply_cltn2 east meridian_reflect = west"
  and "apply_cltn2 west meridian_reflect = east"
  and "apply_cltn2 north meridian_reflect = north"
  and "apply_cltn2 south meridian_reflect = south"
  and "apply_cltn2 K2_centre meridian_reflect = K2_centre"
  and "apply_cltn2 east equator_reflect = east"
  and "apply_cltn2 west equator_reflect = west"
  and "apply_cltn2 north equator_reflect = south"
  and "apply_cltn2 south equator_reflect = north"
  and "apply_cltn2 K2_centre equator_reflect = K2_centre"
proof -
  have "(vector [1,0,1] :: real^3) v* rep_meridian_reflect = vector [-1,0,1]"
    and "(vector [-1,0,1] :: real^3) v* rep_meridian_reflect = vector [1,0,1]"
    and "(vector [0,1,1] :: real^3) v* rep_meridian_reflect = vector [0,1,1]"
    and "(vector [0,-1,1] :: real^3) v* rep_meridian_reflect = vector [0,-1,1]"
    and "(vector [0,0,1] :: real^3) v* rep_meridian_reflect = vector [0,0,1]"
    and "(vector [1,0,1] :: real^3) v* rep_equator_reflect = vector [1,0,1]"
    and "(vector [-1,0,1] :: real^3) v* rep_equator_reflect = vector [-1,0,1]"
    and "(vector [0,1,1] :: real^3) v* rep_equator_reflect = vector [0,-1,1]"
    and "(vector [0,-1,1] :: real^3) v* rep_equator_reflect = vector [0,1,1]"
    and "(vector [0,0,1] :: real^3) v* rep_equator_reflect = vector [0,0,1]"
    unfolding rep_meridian_reflect_def and rep_equator_reflect_def
      and vector_matrix_mult_def
    by (simp_all add: vec_eq_iff forall_3 vector_3 sum_3)
  with compass_reflect_invertible and compass_non_zero and K2_centre_non_zero
  show "apply_cltn2 east meridian_reflect = west"
    and "apply_cltn2 west meridian_reflect = east"
    and "apply_cltn2 north meridian_reflect = north"
    and "apply_cltn2 south meridian_reflect = south"
    and "apply_cltn2 K2_centre meridian_reflect = K2_centre"
    and "apply_cltn2 east equator_reflect = east"
    and "apply_cltn2 west equator_reflect = west"
    and "apply_cltn2 north equator_reflect = south"
    and "apply_cltn2 south equator_reflect = north"
    and "apply_cltn2 K2_centre equator_reflect = K2_centre"
    unfolding compass_defs and K2_centre_def
      and meridian_reflect_def and equator_reflect_def
    by (simp_all add: apply_cltn2_abs)
qed

lemma on_equator_rep:
  assumes "z_non_zero a" and "proj2_incident a equator"
  shows " x. a = proj2_abs (vector [x,0,1])"
proof -
  let ?ra = "proj2_rep a"
  let ?ca1 = "cart2_append1 a"
  let ?x = "?ca1$1"
  from compass_non_zero and proj2_incident a equator
  have "?ra  vector [0,1,0] = 0"
    by (unfold equator_def) (simp add: proj2_incident_right_abs)
  hence "?ra$2 = 0" by (unfold inner_vec_def vector_def) (simp add: sum_3)
  hence "?ca1$2 = 0" by (unfold cart2_append1_def) simp
  moreover
  from z_non_zero a have "?ca1$3 = 1" by (rule cart2_append1_z)
  ultimately
  have "?ca1 = vector [?x,0,1]"
    by (unfold vector_def) (simp add: vec_eq_iff forall_3)
  with z_non_zero a
  have "proj2_abs (vector [?x,0,1]) = a" by (simp add: proj2_abs_cart2_append1)
  thus " x. a = proj2_abs (vector [x,0,1])" by (simp add: exI [of _ ?x])
qed

lemma on_meridian_rep:
  assumes "z_non_zero a" and "proj2_incident a meridian"
  shows " y. a = proj2_abs (vector [0,y,1])"
proof -
  let ?ra = "proj2_rep a"
  let ?ca1 = "cart2_append1 a"
  let ?y = "?ca1$2"
  from compass_non_zero and proj2_incident a meridian
  have "?ra  vector [1,0,0] = 0"
    by (unfold meridian_def) (simp add: proj2_incident_right_abs)
  hence "?ra$1 = 0" by (unfold inner_vec_def vector_def) (simp add: sum_3)
  hence "?ca1$1 = 0" by (unfold cart2_append1_def) simp
  moreover
  from z_non_zero a have "?ca1$3 = 1" by (rule cart2_append1_z)
  ultimately
  have "?ca1 = vector [0,?y,1]"
    by (unfold vector_def) (simp add: vec_eq_iff forall_3)
  with z_non_zero a
  have "proj2_abs (vector [0,?y,1]) = a" by (simp add: proj2_abs_cart2_append1)
  thus " y. a = proj2_abs (vector [0,y,1])" by (simp add: exI [of _ ?y])
qed

subsection ‹Definition of the Klein--Beltrami model of the hyperbolic plane›
abbreviation "hyp2 == K2"

typedef hyp2 = K2
  using K2_centre_in_K2
  by auto

definition hyp2_rep :: "hyp2  real^2" where
  "hyp2_rep p  cart2_pt (Rep_hyp2 p)"

definition hyp2_abs :: "real^2  hyp2" where
  "hyp2_abs v = Abs_hyp2 (proj2_pt v)"

lemma norm_lt_1_iff_in_hyp2:
  shows "norm v < 1  proj2_pt v  hyp2"
proof -
  let ?v' = "vector2_append1 v"
  have "?v'  0" by (rule vector2_append1_non_zero)

  from real_less_rsqrt [of "norm v" 1]
    and abs_square_less_1 [of "norm v"]
  have "norm v < 1  (norm v)2 < 1" by auto
  hence "norm v < 1  ?v'  (M *v ?v') < 0" by (simp add: norm_M)
  with ?v'  0 have "norm v < 1  proj2_abs ?v'  K2" by (subst K2_abs)
  thus "norm v < 1  proj2_pt v  hyp2" by (unfold proj2_pt_def)
qed

lemma norm_eq_1_iff_in_S:
  shows "norm v = 1  proj2_pt v  S"
proof -
  let ?v' = "vector2_append1 v"
  have "?v'  0" by (rule vector2_append1_non_zero)

  from real_sqrt_unique [of "norm v" 1]
  have "norm v = 1  (norm v)2 = 1" by auto
  hence "norm v = 1  ?v'  (M *v ?v') = 0" by (simp add: norm_M)
  with ?v'  0 have "norm v = 1  proj2_abs ?v'  S" by (subst S_abs)
  thus "norm v = 1  proj2_pt v  S" by (unfold proj2_pt_def)
qed

lemma norm_le_1_iff_in_hyp2_S:
  "norm v  1  proj2_pt v  hyp2  S"
  using norm_lt_1_iff_in_hyp2 [of v] and norm_eq_1_iff_in_S [of v]
  by auto

lemma proj2_pt_hyp2_rep: "proj2_pt (hyp2_rep p) = Rep_hyp2 p"
proof -
  let ?p' = "Rep_hyp2 p"
  let ?v = "proj2_rep ?p'"
  have "?v  0" by (rule proj2_rep_non_zero)

  have "proj2_abs ?v = ?p'" by (rule proj2_abs_rep)

  have "?p'  hyp2" by (rule Rep_hyp2)
  with ?v  0 and proj2_abs ?v = ?p'
  have "?v  (M *v ?v) < 0" by (simp add: K2_imp_M_neg)
  hence "?v$3  0" by (rule M_neg_imp_z_non_zero)
  hence "proj2_pt (cart2_pt ?p') = ?p'" by (rule proj2_cart2)
  thus "proj2_pt (hyp2_rep p) = ?p'" by (unfold hyp2_rep_def)
qed

lemma hyp2_rep_abs:
  assumes "norm v < 1"
  shows "hyp2_rep (hyp2_abs v) = v"
proof -
  from norm v < 1
  have "proj2_pt v  hyp2" by (simp add: norm_lt_1_iff_in_hyp2)
  hence "Rep_hyp2 (Abs_hyp2 (proj2_pt v)) = proj2_pt v"
    by (simp add: Abs_hyp2_inverse)
  hence "hyp2_rep (hyp2_abs v) = cart2_pt (proj2_pt v)"
    by (unfold hyp2_rep_def hyp2_abs_def) simp
  thus "hyp2_rep (hyp2_abs v) = v" by (simp add: cart2_proj2)
qed

lemma hyp2_abs_rep: "hyp2_abs (hyp2_rep p) = p"
  by (unfold hyp2_abs_def) (simp add: proj2_pt_hyp2_rep Rep_hyp2_inverse)

lemma norm_hyp2_rep_lt_1: "norm (hyp2_rep p) < 1"
proof -
  have "proj2_pt (hyp2_rep p) = Rep_hyp2 p" by (rule proj2_pt_hyp2_rep)
  hence "proj2_pt (hyp2_rep p)  hyp2" by (simp add: Rep_hyp2)
  thus "norm (hyp2_rep p) < 1" by (simp add: norm_lt_1_iff_in_hyp2)
qed

lemma hyp2_S_z_non_zero:
  assumes "p  hyp2  S"
  shows "z_non_zero p"
proof -
  from p  hyp2  S
  have "conic_sgn p  0" by (unfold K2_def S_def) auto
  hence "conic_sgn p  1" by simp
  thus "z_non_zero p" by (rule conic_sgn_not_1_z_non_zero)
qed

lemma hyp2_S_not_equal:
  assumes "a  hyp2" and "p  S"
  shows "a  p"
  using assms and S_K2_empty
  by auto

lemma hyp2_S_cart2_inj:
  assumes "p  hyp2  S" and "q  hyp2  S" and "cart2_pt p = cart2_pt q"
  shows "p = q"
proof -
  from p  hyp2  S and q  hyp2  S
  have "z_non_zero p" and "z_non_zero q" by (simp_all add: hyp2_S_z_non_zero)
  hence "proj2_pt (cart2_pt p) = p" and "proj2_pt (cart2_pt q) = q"
    by (simp_all add: proj2_cart2)

  from cart2_pt p = cart2_pt q
  have "proj2_pt (cart2_pt p) = proj2_pt (cart2_pt q)" by simp
  with proj2_pt (cart2_pt p) = p [symmetric] and proj2_pt (cart2_pt q) = q
  show "p = q" by simp
qed

lemma on_equator_in_hyp2_rep:
  assumes "a  hyp2" and "proj2_incident a equator"
  shows " x. ¦x¦ < 1  a = proj2_abs (vector [x,0,1])"
proof -
  from a  hyp2 have "z_non_zero a" by (simp add: hyp2_S_z_non_zero)
  with proj2_incident a equator and on_equator_rep
  obtain x where "a = proj2_abs (vector [x,0,1])" (is "a = proj2_abs ?v")
    by auto

  have "?v  0" by (simp add: vec_eq_iff forall_3 vector_3)
  with a  hyp2 and a = proj2_abs ?v
  have "?v  (M *v ?v) < 0" by (simp add: K2_abs)
  hence "x2 < 1"
    unfolding M_def matrix_vector_mult_def inner_vec_def
    by (simp add: sum_3 vector_3 power2_eq_square)
  with real_sqrt_abs [of x] and real_sqrt_less_iff [of "x2" 1]
  have "¦x¦ < 1" by simp
  with a = proj2_abs ?v
  show " x. ¦x¦ < 1  a = proj2_abs (vector [x,0,1])"
    by (simp add: exI [of _ x])
qed

lemma on_meridian_in_hyp2_rep:
  assumes "a  hyp2" and "proj2_incident a meridian"
  shows " y. ¦y¦ < 1  a = proj2_abs (vector [0,y,1])"
proof -
  from a  hyp2 have "z_non_zero a" by (simp add: hyp2_S_z_non_zero)
  with proj2_incident a meridian and on_meridian_rep
  obtain y where "a = proj2_abs (vector [0,y,1])" (is "a = proj2_abs ?v")
    by auto

  have "?v  0" by (simp add: vec_eq_iff forall_3 vector_3)
  with a  hyp2 and a = proj2_abs ?v
  have "?v  (M *v ?v) < 0" by (simp add: K2_abs)
  hence "y2 < 1"
    unfolding M_def matrix_vector_mult_def inner_vec_def
    by (simp add: sum_3 vector_3 power2_eq_square)
  with real_sqrt_abs [of y] and real_sqrt_less_iff [of "y2" 1]
  have "¦y¦ < 1" by simp
  with a = proj2_abs ?v
  show " y. ¦y¦ < 1  a = proj2_abs (vector [0,y,1])"
    by (simp add: exI [of _ y])
qed

definition hyp2_cltn2 :: "hyp2  cltn2  hyp2" where
  "hyp2_cltn2 p A  Abs_hyp2 (apply_cltn2 (Rep_hyp2 p) A)"

definition is_K2_isometry :: "cltn2  bool" where
  "is_K2_isometry J  ( p. apply_cltn2 p J  S  p  S)"

lemma cltn2_id_is_K2_isometry: "is_K2_isometry cltn2_id"
  unfolding is_K2_isometry_def
  by simp

lemma J_M_J_transpose_K2_isometry:
  assumes "k  0"
  and "repJ ** M ** transpose repJ = k *R M" (is "?N = _")
  shows "is_K2_isometry (cltn2_abs repJ)" (is "is_K2_isometry ?J")
proof -
  from ?N = k *R M
  have "?N ** ((1/k) *R M) = mat 1"
    by (simp add: matrix_scalar_ac k  0 M_self_inverse)
  with right_invertible_iff_invertible [of repJ]
  have "invertible repJ"
    by (simp add: matrix_mul_assoc
      exI [of _ "M ** transpose repJ ** ((1/k) *R M)"])

  have " t. apply_cltn2 t ?J  S  t  S"
  proof
    fix t :: proj2
    have "proj2_rep t  ((k *R M) *v proj2_rep t)
      = k * (proj2_rep t  (M *v proj2_rep t))"
      by (simp add: scaleR_matrix_vector_assoc [symmetric]  dot_scaleR_mult)
    with ?N = k *R M
    have "proj2_rep t  (?N *v proj2_rep t)
      = k * (proj2_rep t  (M *v proj2_rep t))"
      by simp
    hence "proj2_rep t  (?N *v proj2_rep t) = 0
       k * (proj2_rep t  (M *v proj2_rep t)) = 0"
      by simp
    with k  0
    have "proj2_rep t  (?N *v proj2_rep t) = 0
       proj2_rep t  (M *v proj2_rep t) = 0"
      by simp
    with invertible repJ
    have "apply_cltn2 t ?J  S  proj2_rep t  (M *v proj2_rep t) = 0"
      by (simp add: apply_cltn2_right_abs_in_S)
    thus "apply_cltn2 t ?J  S  t  S" by (unfold S_alt_def)
  qed
  thus "is_K2_isometry ?J" by (unfold is_K2_isometry_def)
qed

lemma equator_reflect_K2_isometry:
  shows "is_K2_isometry equator_reflect"
  unfolding compass_reflect_defs
  by (rule J_M_J_transpose_K2_isometry [of 1])
     (simp_all add: M_def matrix_matrix_mult_def transpose_def
       vec_eq_iff forall_3 sum_3 vector_3)

lemma meridian_reflect_K2_isometry:
  shows "is_K2_isometry meridian_reflect"
  unfolding compass_reflect_defs
  by (rule J_M_J_transpose_K2_isometry [of 1])
     (simp_all add: M_def matrix_matrix_mult_def transpose_def
       vec_eq_iff forall_3 sum_3 vector_3)

lemma cltn2_compose_is_K2_isometry:
  assumes "is_K2_isometry H" and "is_K2_isometry J"
  shows "is_K2_isometry (cltn2_compose H J)"
  using is_K2_isometry H and is_K2_isometry J
  unfolding is_K2_isometry_def
  by (simp add: cltn2.act_act [simplified, symmetric])

lemma cltn2_inverse_is_K2_isometry:
  assumes "is_K2_isometry J"
  shows "is_K2_isometry (cltn2_inverse J)"
proof -
  { fix p
    from is_K2_isometry J
    have "apply_cltn2 p (cltn2_inverse J)  S
       apply_cltn2 (apply_cltn2 p (cltn2_inverse J)) J  S"
      unfolding is_K2_isometry_def
      by simp
    hence "apply_cltn2 p (cltn2_inverse J)  S  p  S"
      by (simp add: cltn2.act_inv_act [simplified]) }
  thus "is_K2_isometry (cltn2_inverse J)"
    unfolding is_K2_isometry_def ..
qed

interpretation K2_isometry_subgroup: subgroup
  "Collect is_K2_isometry"
  "(|carrier = UNIV, mult = cltn2_compose, one = cltn2_id|)"
  unfolding subgroup_def
  by (simp add:
    cltn2_id_is_K2_isometry
    cltn2_compose_is_K2_isometry
    cltn2_inverse_is_K2_isometry)

interpretation K2_isometry: group
  "(|carrier = Collect is_K2_isometry, mult = cltn2_compose, one = cltn2_id|)"
  using cltn2.is_group and K2_isometry_subgroup.subgroup_is_group
  by simp

lemma K2_isometry_inverse_inv [simp]:
  assumes "is_K2_isometry J"
  shows "inv(|carrier = Collect is_K2_isometry, mult = cltn2_compose, one = cltn2_id|)J
  = cltn2_inverse J"
  using cltn2_left_inverse
    and is_K2_isometry J
    and cltn2_inverse_is_K2_isometry
    and K2_isometry.inv_equality
  by simp

definition real_hyp2_C :: "[hyp2, hyp2, hyp2, hyp2]  bool"
  ("_ _ \<congruent>K _ _" [99,99,99,99] 50) where
  "p q \<congruent>K r s 
    ( A. is_K2_isometry A  hyp2_cltn2 p A = r  hyp2_cltn2 q A = s)"

definition real_hyp2_B :: "[hyp2, hyp2, hyp2]  bool"
("BK _ _ _" [99,99,99] 50) where
  "BK p q r  B (hyp2_rep p) (hyp2_rep q) (hyp2_rep r)"


subsection ‹$K$-isometries map the interior of the conic to itself›

lemma collinear_quadratic:
  assumes "t = i *R a + r"
  shows "t  (M *v t) =
  (a  (M *v a)) * i2 + 2 * (a  (M *v r)) * i + r  (M *v r)"
proof -
  from M_reverse have "i * (a  (M *v r)) = i * (r  (M *v a))" by simp
  with t = i *R a + r
  show "t  (M *v t) =
    (a  (M *v a)) * i2 + 2 * (a  (M *v r)) * i + r  (M *v r)"
    by (simp add:
      inner_add_left
      matrix_vector_right_distrib
      inner_add_right
      matrix_scaleR_vector_ac
      inner_scaleR_right
      scaleR_matrix_vector_assoc [symmetric]
      M_reverse
      power2_eq_square
      algebra_simps)
qed

lemma S_quadratic':
  assumes "p  0" and "q  0" and "proj2_abs p  proj2_abs q"
  shows "proj2_abs (k *R p + q)  S
   p  (M *v p) * k2 + p  (M *v q) * 2 * k + q  (M *v q) = 0"
proof -
  let ?r = "k *R p + q"
  from p  0 and q  0 and proj2_abs p  proj2_abs q
    and dependent_proj2_abs [of p q k 1]
  have "?r  0" by auto
  hence "proj2_abs ?r  S  ?r  (M *v ?r) = 0" by (rule S_abs)
  with collinear_quadratic [of ?r k p q]
  show "proj2_abs ?r  S
     p  (M *v p) * k2 + p  (M *v q) * 2 * k + q  (M *v q) = 0"
    by (simp add: dot_lmul_matrix [symmetric] algebra_simps)
qed

lemma S_quadratic:
  assumes "p  q" and "r = proj2_abs (k *R proj2_rep p + proj2_rep q)"
  shows "r  S
   proj2_rep p  (M *v proj2_rep p) * k2
      + proj2_rep p  (M *v proj2_rep q) * 2 * k
      + proj2_rep q  (M *v proj2_rep q)
    = 0"
proof -
  let ?u = "proj2_rep p"
  let ?v = "proj2_rep q"
  let ?w = "k *R ?u + ?v"
  have "?u  0" and "?v  0" by (rule proj2_rep_non_zero)+

  from p  q have "proj2_abs ?u  proj2_abs ?v" by (simp add: proj2_abs_rep)
  with ?u  0 and ?v  0 and r = proj2_abs ?w
  show "r  S
     ?u  (M *v ?u) * k2 + ?u  (M *v ?v) * 2 * k + ?v  (M *v ?v) = 0"
    by (simp add: S_quadratic')
qed

definition quarter_discrim :: "real^3  real^3  real" where
  "quarter_discrim p q  (p  (M *v q))2 - p  (M *v p) * (q  (M *v q))"

lemma quarter_discrim_invariant:
  assumes "t = i *R a + r"
  shows "quarter_discrim a t = quarter_discrim a r"
proof -
  from t = i *R a + r
  have "a  (M *v t) = i * (a  (M *v a)) + a  (M *v r)"
    by (simp add:
      matrix_vector_right_distrib
      inner_add_right
      matrix_scaleR_vector_ac
      scaleR_matrix_vector_assoc [symmetric])
  hence "(a  (M *v t))2 =
    (a  (M *v a))2 * i2 +
    2 * (a  (M *v a)) * (a  (M *v r)) * i +
    (a  (M *v r))2"
    by (simp add: power2_eq_square algebra_simps)
  moreover from collinear_quadratic and t = i *R a + r
  have "a  (M *v a) * (t  (M *v t)) =
    (a  (M *v a))2 * i2 +
    2 * (a  (M *v a)) * (a  (M *v r)) * i +
    a  (M *v a) * (r  (M *v r))"
    by (simp add: power2_eq_square algebra_simps)
  ultimately show "quarter_discrim a t = quarter_discrim a r"
    by (unfold quarter_discrim_def, simp)
qed

lemma quarter_discrim_positive:
  assumes "p  0"  and "q  0" and "proj2_abs p  proj2_abs q" (is "?pp  ?pq")
  and "proj2_abs p  K2"
  shows "quarter_discrim p q > 0"
proof -
  let ?i = "-q$3/p$3"
  let ?t = "?i *R p + q"

  from p  0 and ?pp  K2
  have "p  (M *v p) < 0" by (subst K2_abs [symmetric])
  hence "p$3  0" by (rule M_neg_imp_z_non_zero)
  hence "?t$3 = 0" by simp
  hence "?t  (M *v ?t) = (?t$1)2 + (?t$2)2"
    unfolding matrix_vector_mult_def and M_def and vector_def
    by (simp add: inner_vec_def sum_3 power2_eq_square)

  from p$3  0 have "p  0" by auto
  with q  0 and ?pp  ?pq and dependent_proj2_abs [of p q ?i 1]
  have "?t  0" by auto
  with ?t$3 = 0 have "?t$1  0  ?t$2  0" by (simp add: vec_eq_iff forall_3)
  hence "(?t$1)2 > 0  (?t$2)2 > 0" by simp
  moreover have "(?t$2)2  0" and "(?t$1)2  0" by simp_all
  ultimately have "(?t$1)2 + (?t$2)2 > 0" by arith
  with ?t  (M *v ?t) = (?t$1)2 + (?t$2)2 have "?t  (M *v ?t) > 0" by simp
  with mult_neg_pos [of "p  (M *v p)"] and p  (M *v p) < 0
  have "p  (M *v p) * (?t  (M *v ?t)) < 0" by simp
  moreover have "(p  (M *v ?t))2  0" by simp
  ultimately
  have "(p  (M *v ?t))2 - p  (M *v p) * (?t  (M *v ?t)) > 0" by arith
  with quarter_discrim_invariant [of ?t ?i p q]
  show "quarter_discrim p q > 0" by (unfold quarter_discrim_def, simp)
qed

lemma quarter_discrim_self_zero:
  assumes "proj2_abs a = proj2_abs b"
  shows "quarter_discrim a b = 0"
proof cases
  assume "b = 0"
  thus "quarter_discrim a b = 0" by (unfold quarter_discrim_def, simp)
next
  assume "b  0"
  with proj2_abs a = proj2_abs b and proj2_abs_abs_mult
  obtain k where "a = k *R b" by auto
  thus "quarter_discrim a b = 0"
    unfolding quarter_discrim_def
    by (simp add: power2_eq_square
      matrix_scaleR_vector_ac
      scaleR_matrix_vector_assoc [symmetric])
qed

definition S_intersection_coeff1 :: "real^3  real^3  real" where
  "S_intersection_coeff1 p q
   (-p  (M *v q) + sqrt (quarter_discrim p q)) / (p  (M *v p))"

definition S_intersection_coeff2 :: "real^3  real^3  real" where
  "S_intersection_coeff2 p q
   (-p  (M *v q) - sqrt (quarter_discrim p q)) / (p  (M *v p))"

definition S_intersection1_rep :: "real^3  real^3  real^3" where
  "S_intersection1_rep p q  (S_intersection_coeff1 p q) *R p + q"

definition S_intersection2_rep :: "real^3  real^3  real^3" where
  "S_intersection2_rep p q  (S_intersection_coeff2 p q) *R p + q"

definition S_intersection1 :: "real^3  real^3  proj2" where
  "S_intersection1 p q  proj2_abs (S_intersection1_rep p q)"

definition S_intersection2 :: "real^3  real^3  proj2" where
  "S_intersection2 p q  proj2_abs (S_intersection2_rep p q)"

lemmas S_intersection_coeffs_defs =
  S_intersection_coeff1_def S_intersection_coeff2_def

lemmas S_intersections_defs =
  S_intersection1_def S_intersection2_def
  S_intersection1_rep_def S_intersection2_rep_def

lemma S_intersection_coeffs_distinct:
  assumes "p  0" and "q  0" and "proj2_abs p  proj2_abs q" (is "?pp  ?pq")
  and "proj2_abs p  K2"
  shows "S_intersection_coeff1 p q  S_intersection_coeff2 p q"
proof -
  from p  0 and ?pp  K2
  have "p  (M *v p) < 0" by (subst K2_abs [symmetric])

  from assms have "quarter_discrim p q > 0" by (rule quarter_discrim_positive)
  with p  (M *v p) < 0
  show "S_intersection_coeff1 p q  S_intersection_coeff2 p q"
    by (unfold S_intersection_coeffs_defs, simp)
qed

lemma S_intersections_distinct:
  assumes "p  0" and "q  0" and "proj2_abs p  proj2_abs q" (is "?pp  ?pq")
  and "proj2_abs p  K2"
  shows "S_intersection1 p q  S_intersection2 p q"
proof-
  from p  0 and q  0 and ?pp  ?pq and ?pp  K2
  have "S_intersection_coeff1 p q  S_intersection_coeff2 p q"
    by (rule S_intersection_coeffs_distinct)
  with p  0 and q  0 and ?pp  ?pq and proj2_Col_coeff_unique'
  show "S_intersection1 p q  S_intersection2 p q"
    by (unfold S_intersections_defs, auto)
qed

lemma S_intersections_in_S:
  assumes "p  0"  and "q  0" and "proj2_abs p  proj2_abs q" (is "?pp  ?pq")
  and "proj2_abs p  K2"
  shows "S_intersection1 p q  S" and "S_intersection2 p q  S"
proof -
  let ?j = "S_intersection_coeff1 p q"
  let ?k = "S_intersection_coeff2 p q"
  let ?a = "p  (M *v p)"
  let ?b = "2 * (p  (M *v q))"
  let ?c = "q  (M *v q)"

  from p  0 and ?pp  K2 have "?a < 0" by (subst K2_abs [symmetric])

  have qd: "discrim ?a ?b ?c = 4 * quarter_discrim p q"
    unfolding discrim_def quarter_discrim_def
    by (simp add: power2_eq_square)
  with times_divide_times_eq [of
    2 2 "sqrt (quarter_discrim p q) - p  (M *v q)" ?a]
    and times_divide_times_eq [of
    2 2 "-p  (M *v q) - sqrt (quarter_discrim p q)" ?a]
    and real_sqrt_mult and real_sqrt_abs [of 2]
  have "?j = (-?b + sqrt (discrim ?a ?b ?c)) / (2 * ?a)"
    and "?k = (-?b - sqrt (discrim ?a ?b ?c)) / (2 * ?a)"
    by (unfold S_intersection_coeffs_defs, simp_all add: algebra_simps)

  from assms have "quarter_discrim p q > 0" by (rule quarter_discrim_positive)
  with qd
  have "discrim (p  (M *v p)) (2 * (p  (M *v q))) (q  (M *v q)) > 0"
    by simp
  with ?j = (-?b + sqrt (discrim ?a ?b ?c)) / (2 * ?a)
    and ?k = (-?b - sqrt (discrim ?a ?b ?c)) / (2 * ?a)
    and ?a < 0 and discriminant_nonneg [of ?a ?b ?c ?j]
    and discriminant_nonneg [of ?a ?b ?c ?k]
  have "p  (M *v p) * ?j2 + 2 * (p  (M *v q)) * ?j + q  (M *v q) = 0"
    and "p  (M *v p) * ?k2 + 2 * (p  (M *v q)) * ?k + q  (M *v q) = 0"
    by (unfold S_intersection_coeffs_defs, auto)
  with p  0 and q  0 and ?pp  ?pq and S_quadratic'
  show "S_intersection1 p q  S" and "S_intersection2 p q  S"
    by (unfold S_intersections_defs, simp_all)
qed

lemma S_intersections_Col:
  assumes "p  0" and "q  0"
  shows "proj2_Col (proj2_abs p) (proj2_abs q) (S_intersection1 p q)"
  (is "proj2_Col ?pp ?pq ?pr")
    and "proj2_Col (proj2_abs p) (proj2_abs q) (S_intersection2 p q)"
  (is "proj2_Col ?pp ?pq ?ps")
proof -
  { assume "?pp = ?pq"
    hence "proj2_Col ?pp ?pq ?pr" and "proj2_Col ?pp ?pq ?ps"
      by (simp_all add: proj2_Col_coincide) }
  moreover
  { assume "?pp  ?pq"
    with p  0 and q  0 and dependent_proj2_abs [of p q _ 1]
    have "S_intersection1_rep p q  0" (is "?r  0")
      and "S_intersection2_rep p q  0" (is "?s  0")
      by (unfold S_intersection1_rep_def S_intersection2_rep_def, auto)
    with p  0 and q  0
      and proj2_Col_abs [of p q ?r "S_intersection_coeff1 p q" 1 "-1"]
      and proj2_Col_abs [of p q ?s "S_intersection_coeff2 p q" 1 "-1"]
    have "proj2_Col ?pp ?pq ?pr" and "proj2_Col ?pp ?pq ?ps"
      by (unfold S_intersections_defs, simp_all) }
  ultimately show "proj2_Col ?pp ?pq ?pr" and "proj2_Col ?pp ?pq ?ps" by fast+
qed

lemma S_intersections_incident:
  assumes "p  0" and "q  0" and "proj2_abs p  proj2_abs q" (is "?pp  ?pq")
  and "proj2_incident (proj2_abs p) l" and "proj2_incident (proj2_abs q) l"
  shows "proj2_incident (S_intersection1 p q) l" (is "proj2_incident ?pr l")
  and "proj2_incident (S_intersection2 p q) l" (is "proj2_incident ?ps l")
proof -
  from p  0 and q  0
  have "proj2_Col ?pp ?pq ?pr" and "proj2_Col ?pp ?pq ?ps"
    by (rule S_intersections_Col)+
  with ?pp  ?pq and proj2_incident ?pp l and proj2_incident ?pq l
    and proj2_incident_iff_Col
  show "proj2_incident ?pr l" and "proj2_incident ?ps l" by fast+
qed

lemma K2_line_intersect_twice:
  assumes "a  K2" and "a  r"
  shows " s u. s  u  s  S  u  S  proj2_Col a r s  proj2_Col a r u"
proof -
  let ?a' = "proj2_rep a"
  let ?r' = "proj2_rep r"
  from proj2_rep_non_zero have "?a'  0" and "?r'  0" by simp_all

  from ?a'  0 and K2_imp_M_neg and proj2_abs_rep and a  K2
  have "?a'  (M *v ?a') < 0" by simp

  from a  r have "proj2_abs ?a'  proj2_abs ?r'" by (simp add: proj2_abs_rep)

  from a  K2 have "proj2_abs ?a'  K2" by (simp add: proj2_abs_rep)
  with ?a'  0 and ?r'  0 and proj2_abs ?a'  proj2_abs ?r'
  have "S_intersection1 ?a' ?r'  S_intersection2 ?a' ?r'" (is "?s  ?u")
    by (rule S_intersections_distinct)

  from ?a'  0 and ?r'  0 and proj2_abs ?a'  proj2_abs ?r'
    and proj2_abs ?a'  K2
  have "?s  S" and "?u  S" by (rule S_intersections_in_S)+

  from ?a'  0 and ?r'  0
  have "proj2_Col (proj2_abs ?a') (proj2_abs ?r') ?s"
    and "proj2_Col (proj2_abs ?a') (proj2_abs ?r') ?u"
    by (rule S_intersections_Col)+
  hence "proj2_Col a r ?s" and "proj2_Col a r ?u"
    by (simp_all add: proj2_abs_rep)
  with ?s  ?u and ?s  S and ?u  S
  show " s u. s  u  s  S  u  S  proj2_Col a r s  proj2_Col a r u"
    by auto
qed

lemma point_in_S_polar_is_tangent:
  assumes "p  S" and "q  S" and "proj2_incident q (polar p)"
  shows "q = p"
proof -
  from p  S have "proj2_incident p (polar p)"
    by (subst incident_own_polar_in_S)

  from line_incident_point_not_in_S
  obtain r where "r  S" and "proj2_incident r (polar p)" by auto
  let ?u = "proj2_rep r"
  let ?v = "proj2_rep p"
  from r  S and p  S and q  S have "r  p" and "q  r" by auto
  with proj2_incident p (polar p)
    and proj2_incident q (polar p)
    and proj2_incident r (polar p)
    and proj2_incident_iff [of r p "polar p" q]
  obtain k where "q = proj2_abs (k *R ?u + ?v)" by auto
  with r  p and q  S and S_quadratic
  have "?u  (M *v ?u) * k2 + ?u  (M *v ?v) * 2 * k + ?v  (M *v ?v) = 0"
    by simp
  moreover from p  S have "?v  (M *v ?v) = 0" by (unfold S_alt_def)
  moreover from proj2_incident r (polar p)
  have "?u  (M *v ?v) = 0" by (unfold incident_polar)
  moreover from r  S have "?u  (M *v ?u)  0" by (unfold S_alt_def)
  ultimately have "k = 0" by simp
  with q = proj2_abs (k *R ?u + ?v)
  show "q = p" by (simp add: proj2_abs_rep)
qed

lemma line_through_K2_intersect_S_twice:
  assumes "p  K2" and "proj2_incident p l"
  shows " q r. q  r  q  S  r  S  proj2_incident q l  proj2_incident r l"
proof -
  from proj2_another_point_on_line
  obtain s where "s  p" and "proj2_incident s l" by auto
  from p  K2 and s  p and K2_line_intersect_twice [of p s]
  obtain q and r where "q  r" and "q  S" and "r  S"
    and "proj2_Col p s q" and "proj2_Col p s r"
    by auto
  with s  p and proj2_incident p l and proj2_incident s l
    and proj2_incident_iff_Col [of p s]
  have "proj2_incident q l" and "proj2_incident r l" by fast+
  with q  r and q  S and r  S
  show " q r. q  r  q  S  r  S  proj2_incident q l  proj2_incident r l"
    by auto
qed

lemma line_through_K2_intersect_S_again:
  assumes "p  K2" and "proj2_incident p l"
  shows " r. r  q  r  S  proj2_incident r l"
proof -
  from p  K2 and proj2_incident p l
    and line_through_K2_intersect_S_twice [of p l]
  obtain s and t where "s  t" and "s  S" and "t  S"
    and "proj2_incident s l" and "proj2_incident t l"
    by auto
  show " r. r  q  r  S  proj2_incident r l"
  proof cases
    assume "t = q"
    with s  t and s  S and proj2_incident s l
    have "s  q  s  S  proj2_incident s l" by simp
    thus " r. r  q  r  S  proj2_incident r l" ..
  next
    assume "t  q"
    with t  S and proj2_incident t l
    have "t  q  t  S  proj2_incident t l" by simp
    thus " r. r  q  r  S  proj2_incident r l" ..
  qed
qed

lemma line_through_K2_intersect_S:
  assumes "p  K2" and "proj2_incident p l"
  shows " r. r  S  proj2_incident r l"
proof -
  from assms
  have " r. r  p  r  S  proj2_incident r l"
    by (rule line_through_K2_intersect_S_again)
  thus " r. r  S  proj2_incident r l" by auto
qed

lemma line_intersect_S_at_most_twice:
  " p q.  rS. proj2_incident r l  r = p  r = q"
proof -
  from line_incident_point_not_in_S
  obtain s where "s  S" and "proj2_incident s l" by auto
  let ?v = "proj2_rep s"
  from proj2_another_point_on_line
  obtain t where "t  s" and "proj2_incident t l" by auto
  let ?w = "proj2_rep t"
  have "?v  0" and "?w  0" by (rule proj2_rep_non_zero)+

  let ?a = "?v  (M *v ?v)"
  let ?b = "2 * (?v  (M *v ?w))"
  let ?c = "?w  (M *v ?w)"
  from s  S have "?a  0"
    unfolding S_def and conic_sgn_def
    by auto
  let ?j = "(-?b + sqrt (discrim ?a ?b ?c)) / (2 * ?a)"
  let ?k = "(-?b - sqrt (discrim ?a ?b ?c)) / (2 * ?a)"
  let ?p = "proj2_abs (?j *R ?v + ?w)"
  let ?q = "proj2_abs (?k *R ?v + ?w)"
  have " rS. proj2_incident r l  r = ?p  r = ?q"
  proof
    fix r
    assume "r  S"
    with s  S have "r  s" by auto
    { assume "proj2_incident r l"
      with t  s and r  s and proj2_incident s l and proj2_incident t l
        and proj2_incident_iff [of s t l r]
      obtain i where "r = proj2_abs (i *R ?v + ?w)" by auto
      with r  S and t  s and S_quadratic
      have "?a * i2 + ?b * i + ?c = 0" by simp
      with ?a  0 and discriminant_iff have "i = ?j  i = ?k" by simp
      with r = proj2_abs (i *R ?v + ?w) have "r = ?p  r = ?q" by auto }
    thus "proj2_incident r l  r = ?p  r = ?q" ..
  qed
  thus " p q.  rS. proj2_incident r l  r = p  r = q" by auto
qed

lemma card_line_intersect_S:
  assumes "T  S" and "proj2_set_Col T"
  shows "card T  2"
proof -
  from proj2_set_Col T
  obtain l where " pT. proj2_incident p l" unfolding proj2_set_Col_def ..
  from line_intersect_S_at_most_twice [of l]
  obtain b and c where " aS. proj2_incident a l  a = b  a = c" by auto
  with  pT. proj2_incident p l and T  S
  have "T  {b,c}" by auto
  hence "card T  card {b,c}" by (simp add: card_mono)
  also from card_suc_ge_insert [of b "{c}"] have "  2" by simp
  finally show "card T  2" .
qed

lemma line_S_two_intersections_only:
  assumes "p  q" and "p  S" and "q  S" and "r  S"
  and "proj2_incident p l" and "proj2_incident q l" and "proj2_incident r l"
  shows "r = p  r = q"
proof -
  from p  q have "card {p,q} = 2" by simp

  from p  S and q  S and r  S have "{r,p,q}  S" by simp_all

  from proj2_incident p l and proj2_incident q l and proj2_incident r l
  have "proj2_set_Col {r,p,q}"
    by (unfold proj2_set_Col_def) (simp add: exI [of _ l])
  with {r,p,q}  S have "card {r,p,q}  2" by (rule card_line_intersect_S)

  show "r = p  r = q"
  proof (rule ccontr)
    assume "¬ (r = p  r = q)"
    hence "r  {p,q}" by simp
    with card {p,q} = 2 and card_insert_disjoint [of "{p,q}" r]
    have "card {r,p,q} = 3" by simp
    with card {r,p,q}  2 show False by simp
  qed
qed

lemma line_through_K2_intersect_S_exactly_twice:
  assumes "p  K2" and "proj2_incident p l"
  shows " q r. q  r  q  S  r  S  proj2_incident q l  proj2_incident r l
   ( sS. proj2_incident s l  s = q  s = r)"
proof -
  from p  K2 and proj2_incident p l
    and line_through_K2_intersect_S_twice [of p l]
  obtain q and r where "q  r" and "q  S" and "r  S"
    and "proj2_incident q l" and "proj2_incident r l"
    by auto
  with line_S_two_intersections_only
  show " q r. q  r  q  S  r  S  proj2_incident q l  proj2_incident r l
     ( sS. proj2_incident s l  s = q  s = r)"
    by blast
qed

lemma tangent_not_through_K2:
  assumes "p  S" and "q  K2"
  shows "¬ proj2_incident q (polar p)"
proof
  assume "proj2_incident q (polar p)"
  with q  K2 and line_through_K2_intersect_S_again [of q "polar p" p]
  obtain r where "r  p" and "r  S" and "proj2_incident r (polar p)" by auto
  from p  S and r  S and proj2_incident r (polar p)
  have "r = p" by (rule point_in_S_polar_is_tangent)
  with r  p show False ..
qed

lemma outside_exists_line_not_intersect_S:
  assumes "conic_sgn p = 1"
  shows " l. proj2_incident p l  ( q. proj2_incident q l  q  S)"
proof -
  let ?r = "proj2_intersection (polar p) z_zero"
  have "proj2_incident ?r (polar p)" and "proj2_incident ?r z_zero"
    by (rule proj2_intersection_incident)+
  from proj2_incident ?r z_zero
  have "conic_sgn ?r = 1" by (rule z_zero_conic_sgn_1)
  with conic_sgn p = 1
  have "proj2_rep p  (M *v proj2_rep p) > 0"
    and "proj2_rep ?r  (M *v proj2_rep ?r) > 0"
    by (unfold conic_sgn_def) (simp_all add: sgn_1_pos)

  from proj2_incident ?r (polar p)
  have "proj2_incident p (polar ?r)" by (rule incident_polar_swap)
  hence "proj2_rep p  (M *v proj2_rep ?r) = 0" by (simp add: incident_polar)

  have "p  ?r"
  proof
    assume "p = ?r"
    with proj2_incident ?r (polar p) have "proj2_incident p (polar p)" by simp
    hence "proj2_rep p  (M *v proj2_rep p) = 0" by (simp add: incident_polar)
    with proj2_rep p  (M *v proj2_rep p) > 0 show False by simp
  qed

  let ?l = "proj2_line_through p ?r"
  have "proj2_incident p ?l" and "proj2_incident ?r ?l"
    by (rule proj2_line_through_incident)+

  have " q. proj2_incident q ?l  q  S"
  proof
    fix q
    show "proj2_incident q ?l  q  S"
    proof
      assume "proj2_incident q ?l"
      with p  ?r and proj2_incident p ?l and proj2_incident ?r ?l
      have "q = p  ( k. q = proj2_abs (k *R proj2_rep p + proj2_rep ?r))"
        by (simp add: proj2_incident_iff [of p ?r ?l q])

      show "q  S"
      proof cases
        assume "q = p"
        with conic_sgn p = 1 show "q  S" by (unfold S_def) simp
      next
        assume "q  p"
        with q = p  ( k. q = proj2_abs (k *R proj2_rep p + proj2_rep ?r))
        obtain k where "q = proj2_abs (k *R proj2_rep p + proj2_rep ?r)"
          by auto
        from proj2_rep p  (M *v proj2_rep p) > 0
        have "proj2_rep p  (M *v proj2_rep p) * k2  0"
          by simp
        with proj2_rep p  (M *v proj2_rep ?r) = 0
          and proj2_rep ?r  (M *v proj2_rep ?r) > 0
        have "proj2_rep p  (M *v proj2_rep p) * k2
          + proj2_rep p  (M *v proj2_rep ?r) * 2 * k
          + proj2_rep ?r  (M *v proj2_rep ?r)
          > 0"
          by simp
        with p  ?r and q = proj2_abs (k *R proj2_rep p + proj2_rep ?r)
        show "q  S" by (simp add: S_quadratic)
      qed
    qed
  qed
  with proj2_incident p ?l
  show " l. proj2_incident p l  ( q. proj2_incident q l  q  S)"
    by (simp add: exI [of _ ?l])
qed

lemma lines_through_intersect_S_twice_in_K2:
  assumes " l. proj2_incident p l
   ( q r. q  r  q  S  r  S  proj2_incident q l  proj2_incident r l)"
  shows "p  K2"
proof (rule ccontr)
  assume "p  K2"
  hence "conic_sgn p  0" by (unfold K2_def) simp

  have "¬ ( l. proj2_incident p l  ( q r.
    q  r  q  S  r  S  proj2_incident q l  proj2_incident r l))"
  proof cases
    assume "conic_sgn p = 0"
    hence "p  S" unfolding S_def ..
    hence "proj2_incident p (polar p)" by (simp add: incident_own_polar_in_S)
    let ?l = "polar p"
    have "¬ ( q r.
      q  r  q  S  r  S  proj2_incident q ?l  proj2_incident r ?l)"
    proof
      assume " q r.
        q  r  q  S  r  S  proj2_incident q ?l  proj2_incident r ?l"
      then obtain q and r where "q  r" and "q  S" and "r  S"
        and "proj2_incident q ?l" and "proj2_incident r ?l"
        by auto
      from p  S and q  S and proj2_incident q ?l
        and r  S and proj2_incident r ?l
      have "q = p" and "r = p" by (simp add: point_in_S_polar_is_tangent)+
      with q  r show False by simp
    qed
    with proj2_incident p ?l
    show "¬ ( l. proj2_incident p l  ( q r.
      q  r  q  S  r  S  proj2_incident q l  proj2_incident r l))"
      by auto
  next
    assume "conic_sgn p  0"
    with conic_sgn p  0 have "conic_sgn p > 0" by simp
    hence "sgn (conic_sgn p) = 1" by simp
    hence "conic_sgn p = 1" by (simp add: sgn_conic_sgn)
    with outside_exists_line_not_intersect_S
    obtain l where "proj2_incident p l" and " q. proj2_incident q l  q  S"
      by auto
    have "¬ ( q r.
      q  r  q  S  r  S  proj2_incident q l  proj2_incident r l)"
    proof
      assume " q r.
        q  r  q  S  r  S  proj2_incident q l  proj2_incident r l"
      then obtain q where "q  S" and "proj2_incident q l" by auto
      from proj2_incident q l and  q. proj2_incident q l  q  S
      have "q  S" by simp
      with q  S show False by simp
    qed
    with proj2_incident p l
    show "¬ ( l. proj2_incident p l  ( q r.
      q  r  q  S  r  S  proj2_incident q l  proj2_incident r l))"
      by auto
  qed
  with  l. proj2_incident p l  ( q r.
    q  r  q  S  r  S  proj2_incident q l  proj2_incident r l)
  show False by simp
qed

lemma line_through_hyp2_pole_not_in_hyp2:
  assumes "a  hyp2" and "proj2_incident a l"
  shows "pole l  hyp2"
proof -
  from assms and line_through_K2_intersect_S
  obtain p where "p  S" and "proj2_incident p l" by auto

  from proj2_incident p l
  have "proj2_incident (pole l) (polar p)" by (rule incident_pole_polar)
  with p  S
  show "pole l  hyp2"
    by (auto simp add: tangent_not_through_K2)
qed

lemma statement60_one_way:
  assumes "is_K2_isometry J" and "p  K2"
  shows "apply_cltn2 p J  K2" (is "?p'  K2")
proof -
  let ?J' = "cltn2_inverse J"

  have " l'. proj2_incident ?p' l'  ( q' r'.
    q'  r'  q'  S  r'  S  proj2_incident q' l'  proj2_incident r' l')"
  proof
    fix l'
    let ?l = "apply_cltn2_line l' ?J'"
    show "proj2_incident ?p' l'  ( q' r'.
      q'  r'  q'  S  r'  S  proj2_incident q' l'  proj2_incident r' l')"
    proof
      assume "proj2_incident ?p' l'"
      hence "proj2_incident p ?l"
        by (simp add: apply_cltn2_incident [of p l' ?J']
          cltn2.inv_inv [simplified])
      with p  K2 and line_through_K2_intersect_S_twice [of p ?l]
      obtain q and r where "q  r" and "q  S" and "r  S"
        and "proj2_incident q ?l" and "proj2_incident r ?l"
        by auto
      let ?q' = "apply_cltn2 q J"
      let ?r' = "apply_cltn2 r J"
      from q  r and apply_cltn2_injective [of q J r] have "?q'  ?r'" by auto

      from q  S and r  S and is_K2_isometry J
      have "?q'  S" and "?r'  S" by (unfold is_K2_isometry_def) simp_all

      from proj2_incident q ?l and proj2_incident r ?l
      have "proj2_incident ?q' l'" and "proj2_incident ?r' l'"
        by (simp_all add: apply_cltn2_incident [of _ l' ?J']
          cltn2.inv_inv [simplified])
      with ?q'  ?r' and ?q'  S and ?r'  S
      show " q' r'.
        q'  r'  q'  S  r'  S  proj2_incident q' l'  proj2_incident r' l'"
        by auto
    qed
  qed
  thus "?p'  K2" by (rule lines_through_intersect_S_twice_in_K2)
qed

lemma is_K2_isometry_hyp2_S:
  assumes "p  hyp2  S" and "is_K2_isometry J"
  shows "apply_cltn2 p J  hyp2  S"
proof cases
  assume "p  hyp2"
  with is_K2_isometry J
  have "apply_cltn2 p J  hyp2" by (rule statement60_one_way)
  thus "apply_cltn2 p J  hyp2  S" ..
next
  assume "p  hyp2"
  with p  hyp2  S have "p  S" by simp
  with is_K2_isometry J
  have "apply_cltn2 p J  S" by (unfold is_K2_isometry_def) simp
  thus "apply_cltn2 p J  hyp2  S" ..
qed

lemma is_K2_isometry_z_non_zero:
  assumes "p  hyp2  S" and "is_K2_isometry J"
  shows "z_non_zero (apply_cltn2 p J)"
proof -
  from p  hyp2  S and is_K2_isometry J
  have "apply_cltn2 p J  hyp2  S" by (rule is_K2_isometry_hyp2_S)
  thus "z_non_zero (apply_cltn2 p J)" by (rule hyp2_S_z_non_zero)
qed

lemma cart2_append1_apply_cltn2:
  assumes "p  hyp2  S" and "is_K2_isometry J"
  shows " k. k  0
   cart2_append1 p v* cltn2_rep J = k *R cart2_append1 (apply_cltn2 p J)"
proof -
  have "cart2_append1 p v* cltn2_rep J
    = (1 / (proj2_rep p)$3) *R (proj2_rep p v* cltn2_rep J)"
    by (unfold cart2_append1_def) (simp add: scaleR_vector_matrix_assoc)

  from p  hyp2  S have "(proj2_rep p)$3  0" by (rule hyp2_S_z_non_zero)

  from apply_cltn2_imp_mult [of p J]
  obtain j where "j  0"
    and "proj2_rep p v* cltn2_rep J = j *R proj2_rep (apply_cltn2 p J)"
    by auto

  from p  hyp2  S and is_K2_isometry J
  have "z_non_zero (apply_cltn2 p J)" by (rule is_K2_isometry_z_non_zero)
  hence "proj2_rep (apply_cltn2 p J)
    = (proj2_rep (apply_cltn2 p J))$3 *R cart2_append1 (apply_cltn2 p J)"
    by (rule proj2_rep_cart2_append1)

  let ?k = "1 / (proj2_rep p)$3 * j * (proj2_rep (apply_cltn2 p J))$3"
  from (proj2_rep p)$3  0 and j  0
    and (proj2_rep (apply_cltn2 p J))$3  0
  have "?k  0" by simp

  from cart2_append1 p v* cltn2_rep J
    = (1 / (proj2_rep p)$3) *R (proj2_rep p v* cltn2_rep J)
    and proj2_rep p v* cltn2_rep J = j *R proj2_rep (apply_cltn2 p J)
  have "cart2_append1 p v* cltn2_rep J
    = (1 / (proj2_rep p)$ 3 * j) *R proj2_rep (apply_cltn2 p J)"
    by simp

  from proj2_rep (apply_cltn2 p J)
    = (proj2_rep (apply_cltn2 p J))$3 *R cart2_append1 (apply_cltn2 p J)
  have "(1 / (proj2_rep p)$3 * j) *R proj2_rep (apply_cltn2 p J)
    = (1 / (proj2_rep p)$3 * j) *R ((proj2_rep (apply_cltn2 p J))$3
    *R cart2_append1 (apply_cltn2 p J))"
    by simp
  with cart2_append1 p v* cltn2_rep J
    = (1 / (proj2_rep p)$ 3 * j) *R proj2_rep (apply_cltn2 p J)
  have "cart2_append1 p v* cltn2_rep J = ?k *R cart2_append1 (apply_cltn2 p J)"
    by simp
  with ?k  0
  show " k. k  0
     cart2_append1 p v* cltn2_rep J = k *R cart2_append1 (apply_cltn2 p J)"
    by (simp add: exI [of _ ?k])
qed

subsection ‹The $K$-isometries form a group action›

lemma hyp2_cltn2_id [simp]: "hyp2_cltn2 p cltn2_id = p"
  by (unfold hyp2_cltn2_def) (simp add: Rep_hyp2_inverse)

lemma apply_cltn2_Rep_hyp2:
  assumes "is_K2_isometry J"
  shows "apply_cltn2 (Rep_hyp2 p) J  hyp2"
proof -
  from is_K2_isometry J and Rep_hyp2 [of p]
  show "apply_cltn2 (Rep_hyp2 p) J  K2" by (rule statement60_one_way)
qed

lemma Rep_hyp2_cltn2:
  assumes "is_K2_isometry J"
  shows "Rep_hyp2 (hyp2_cltn2 p J) = apply_cltn2 (Rep_hyp2 p) J"
proof -
  from is_K2_isometry J
  have "apply_cltn2 (Rep_hyp2 p) J  hyp2" by (rule apply_cltn2_Rep_hyp2)
  thus "Rep_hyp2 (hyp2_cltn2 p J) = apply_cltn2 (Rep_hyp2 p) J"
    by (unfold hyp2_cltn2_def) (rule Abs_hyp2_inverse)
qed

lemma hyp2_cltn2_compose:
  assumes "is_K2_isometry H"
  shows "hyp2_cltn2 (hyp2_cltn2 p H) J = hyp2_cltn2 p (cltn2_compose H J)"
proof -
  from is_K2_isometry H
  have "apply_cltn2 (Rep_hyp2 p) H  hyp2" by (rule apply_cltn2_Rep_hyp2)
  thus "hyp2_cltn2 (hyp2_cltn2 p H) J = hyp2_cltn2 p (cltn2_compose H J)"
    by (unfold hyp2_cltn2_def) (simp add: Abs_hyp2_inverse apply_cltn2_compose)
qed

interpretation K2_isometry: action
  "(|carrier = Collect is_K2_isometry, mult = cltn2_compose, one = cltn2_id|)"
  hyp2_cltn2
proof
  let ?G =
    "(|carrier = Collect is_K2_isometry, mult = cltn2_compose, one = cltn2_id|)"
  fix p
  show "hyp2_cltn2 p 𝟭?G= p"
    by (unfold hyp2_cltn2_def) (simp add: Rep_hyp2_inverse)
  fix H J
  show "H  carrier ?G  J  carrier ?G
     hyp2_cltn2 (hyp2_cltn2 p H) J = hyp2_cltn2 p (H ?GJ)"
    by (simp add: hyp2_cltn2_compose)
qed

subsection ‹The Klein--Beltrami model satisfies Tarski's first three axioms›

lemma three_in_S_tangent_intersection_no_3_Col:
  assumes "p  S" and "q  S" and "r  S"
  and "p  q" and "r  {p,q}"
  shows "proj2_no_3_Col {proj2_intersection (polar p) (polar q),r,p,q}"
  (is "proj2_no_3_Col {?s,r,p,q}")
proof -
  let ?T = "{?s,r,p,q}"

  from p  q have "card {p,q} = 2" by simp
  with r  {p,q} have "card {r,p,q} = 3" by simp

  from p  S and q  S and r  S have "{r,p,q}  S" by simp

  have "proj2_incident ?s (polar p)" and "proj2_incident ?s (polar q)"
    by (rule proj2_intersection_incident)+

  have "?s  S"
  proof
    assume "?s  S"
    with p  S and proj2_incident ?s (polar p)
      and q  S and proj2_incident ?s (polar q)
    have "?s = p" and "?s = q" by (simp_all add: point_in_S_polar_is_tangent)
    hence "p = q" by simp
    with p  q show False ..
  qed
  with {r,p,q}  S have "?s  {r,p,q}" by auto
  with card {r,p,q} = 3 have "card {?s,r,p,q} = 4" by simp

  have " t?T. ¬ proj2_set_Col (?T - {t})"
  proof standard+
    fix t
    assume "t  ?T"
    assume "proj2_set_Col (?T - {t})"
    then obtain l where " a  (?T - {t}). proj2_incident a l"
      unfolding proj2_set_Col_def ..

    from proj2_set_Col (?T - {t})
    have "proj2_set_Col (S  (?T - {t}))"
      by (simp add: proj2_subset_Col [of "(S  (?T - {t}))" "?T - {t}"])
    hence "card (S  (?T - {t}))  2" by (simp add: card_line_intersect_S)

    show False
    proof cases
      assume "t = ?s"
      with ?s  {r,p,q} have "?T - {t} = {r,p,q}" by simp
      with {r,p,q}  S have "S  (?T - {t}) = {r,p,q}" by simp
      with card {r,p,q} = 3 and card (S  (?T - {t}))  2 show False by simp
    next
      assume "t  ?s"
      hence "?s  ?T - {t}" by simp
      with  a  (?T - {t}). proj2_incident a l have "proj2_incident ?s l" ..

      from p  q have "{p,q}  ?T - {t}  {}" by auto
      then obtain d where "d  {p,q}" and "d  ?T - {t}" by auto
      from d  ?T - {t} and  a  (?T - {t}). proj2_incident a l
      have "proj2_incident d l" by simp

      from d  {p,q}
        and proj2_incident ?s (polar p)
        and proj2_incident ?s (polar q)
      have "proj2_incident ?s (polar d)" by auto

      from d  {p,q} and {r,p,q}  S have "d  S" by auto
      hence "proj2_incident d (polar d)" by (unfold incident_own_polar_in_S)

      from d  S and ?s  S have "d  ?s" by auto
      with proj2_incident ?s l
        and proj2_incident d l
        and proj2_incident ?s (polar d)
        and proj2_incident d (polar d)
        and proj2_incident_unique
      have "l = polar d" by auto
      with d  S and point_in_S_polar_is_tangent
      have " aS. proj2_incident a l  a = d" by simp
      with  a  (?T - {t}). proj2_incident a l
      have "S  (?T - {t})  {d}" by auto
      with card_mono [of "{d}"] have "card (S  (?T - {t}))  1" by simp
      hence "card ((S  ?T) - {t})  1" by (simp add: Int_Diff)

      have "S  ?T  insert t ((S  ?T) - {t})" by auto
      with card_suc_ge_insert [of t "(S  ?T) - {t}"]
        and card_mono [of "insert t ((S  ?T) - {t})" "S  ?T"]
      have "card (S  ?T)  card ((S  ?T) - {t}) + 1" by simp
      with card ((S  ?T) - {t})  1 have "card (S  ?T)  2" by simp

      from {r,p,q}  S have "{r,p,q}  S  ?T" by simp
      with card {r,p,q} = 3 and card_mono [of "S  ?T" "{r,p,q}"]
      have "card (S  ?T)  3" by simp
      with card (S  ?T)  2 show False by simp
    qed
  qed
  with card ?T = 4 show "proj2_no_3_Col ?T" unfolding proj2_no_3_Col_def ..
qed

lemma statement65_special_case:
  assumes "p  S" and "q  S" and "r  S" and "p  q" and "r  {p,q}"
  shows " J. is_K2_isometry J
   apply_cltn2 east J = p
   apply_cltn2 west J = q
   apply_cltn2 north J = r
   apply_cltn2 far_north J = proj2_intersection (polar p) (polar q)"
proof -
  let ?s = "proj2_intersection (polar p) (polar q)"
  let ?t = "vector [vector [?s,r,p,q], vector [far_north, north, east, west]]
    :: proj2^4^2"
  have "range (($) (?t$1)) = {?s, r, p, q}"
    unfolding image_def
    by (auto simp add: UNIV_4 vector_4)
  with p  S and q  S and r  S and p  q and r  {p,q}
  have "proj2_no_3_Col (range (($) (?t$1)))"
    by (simp add: three_in_S_tangent_intersection_no_3_Col)
  moreover have "range (($) (?t$2)) = {far_north, north, east, west}"
    unfolding image_def
    by (auto simp add: UNIV_4 vector_4)
   with compass_in_S and east_west_distinct and north_not_east_or_west
     and east_west_tangents_far_north
     and three_in_S_tangent_intersection_no_3_Col [of east west north]
   have "proj2_no_3_Col (range (($) (?t$2)))" by simp
   ultimately have " i. proj2_no_3_Col (range (($) (?t$i)))"
     by (simp add: forall_2)
   hence " J.  j. apply_cltn2 (?t$0$j) J = ?t$1$j"
     by (rule statement53_existence)
   moreover have "0 = (2::2)" by simp
   ultimately obtain J where " j. apply_cltn2 (?t$2$j) J = ?t$1$j" by auto
   hence "apply_cltn2 (?t$2$1) J = ?t$1$1"
     and "apply_cltn2 (?t$2$2) J = ?t$1$2"
     and "apply_cltn2 (?t$2$3) J = ?t$1$3"
     and "apply_cltn2 (?t$2$4) J = ?t$1$4"
     by simp_all
   hence "apply_cltn2 east J = p"
     and "apply_cltn2 west J = q"
     and "apply_cltn2 north J = r"
     and "apply_cltn2 far_north J = ?s"
     by (simp_all add: vector_2 vector_4)
   with compass_non_zero
   have "p = proj2_abs (vector [1,0,1] v* cltn2_rep J)"
     and "q = proj2_abs (vector [-1,0,1] v* cltn2_rep J)"
     and "r = proj2_abs (vector [0,1,1] v* cltn2_rep J)"
     and "?s = proj2_abs (vector [0,1,0] v* cltn2_rep J)"
     unfolding compass_defs and far_north_def
     by (simp_all add: apply_cltn2_left_abs)

   let ?N = "cltn2_rep J ** M ** transpose (cltn2_rep J)"
   from M_symmatrix have "symmatrix ?N" by (rule symmatrix_preserve)
   hence "?N$2$1 = ?N$1$2" and "?N$3$1 = ?N$1$3" and "?N$3$2 = ?N$2$3"
     unfolding symmatrix_def and transpose_def
     by (simp_all add: vec_eq_iff)

   from compass_non_zero and apply_cltn2 east J = p and p  S
     and apply_cltn2_abs_in_S [of "vector [1,0,1]" J]
   have "(vector [1,0,1] :: real^3)  (?N *v vector [1,0,1]) = 0"
     unfolding east_def
     by simp
   hence "?N$1$1 + ?N$1$3 + ?N$3$1 + ?N$3$3 = 0"
     unfolding inner_vec_def and matrix_vector_mult_def
     by (simp add: sum_3 vector_3)
   with ?N$3$1 = ?N$1$3 have "?N$1$1 + 2 * (?N$1$3) + ?N$3$3 = 0" by simp

   from compass_non_zero and apply_cltn2 west J = q and q  S
     and apply_cltn2_abs_in_S [of "vector [-1,0,1]" J]
   have "(vector [-1,0,1] :: real^3)  (?N *v vector [-1,0,1]) = 0"
     unfolding west_def
     by simp
   hence "?N$1$1 - ?N$1$3 - ?N$3$1 + ?N$3$3 = 0"
     unfolding inner_vec_def and matrix_vector_mult_def
     by (simp add: sum_3 vector_3)
   with ?N$3$1 = ?N$1$3 have "?N$1$1 - 2 * (?N$1$3) + ?N$3$3 = 0" by simp
   with ?N$1$1 + 2 * (?N$1$3) + ?N$3$3 = 0
   have "?N$1$1 + 2 * (?N$1$3) + ?N$3$3 = ?N$1$1 - 2 * (?N$1$3) + ?N$3$3"
     by simp
   hence "?N$1$3 = 0" by simp
   with ?N$1$1 + 2 * (?N$1$3) + ?N$3$3 = 0 have "?N$3$3 = - (?N$1$1)" by simp

   from compass_non_zero and apply_cltn2 north J = r and r  S
     and apply_cltn2_abs_in_S [of "vector [0,1,1]" J]
   have "(vector [0,1,1] :: real^3)  (?N *v vector [0,1,1]) = 0"
     unfolding north_def
     by simp
   hence "?N$2$2 + ?N$2$3 + ?N$3$2 + ?N$3$3 = 0"
     unfolding inner_vec_def and matrix_vector_mult_def
     by (simp add: sum_3 vector_3)
   with ?N$3$2 = ?N$2$3 have "?N$2$2 + 2 * (?N$2$3) + ?N$3$3 = 0" by simp

   have "proj2_incident ?s (polar p)" and "proj2_incident ?s (polar q)"
     by (rule proj2_intersection_incident)+

   from compass_non_zero
   have "vector [1,0,1] v* cltn2_rep J  0"
     and "vector [-1,0,1] v* cltn2_rep J  0"
     and "vector [0,1,0] v* cltn2_rep J  0"
     by (simp_all add: non_zero_mult_rep_non_zero)
   from vector [1,0,1] v* cltn2_rep J  0
     and vector [-1,0,1] v* cltn2_rep J  0
     and p = proj2_abs (vector [1,0,1] v* cltn2_rep J)
     and q = proj2_abs (vector [-1,0,1] v* cltn2_rep J)
   have "polar p = proj2_line_abs (M *v (vector [1,0,1] v* cltn2_rep J))"
     and "polar q = proj2_line_abs (M *v (vector [-1,0,1] v* cltn2_rep J))"
     by (simp_all add: polar_abs)

   from vector [1,0,1] v* cltn2_rep J  0
     and vector [-1,0,1] v* cltn2_rep J  0
     and M_invertible
   have "M *v (vector [1,0,1] v* cltn2_rep J)  0"
     and "M *v (vector [-1,0,1] v* cltn2_rep J)  0"
     by (simp_all add: invertible_times_non_zero)
   with vector [0,1,0] v* cltn2_rep J  0
     and polar p = proj2_line_abs (M *v (vector [1,0,1] v* cltn2_rep J))
     and polar q = proj2_line_abs (M *v (vector [-1,0,1] v* cltn2_rep J))
     and ?s = proj2_abs (vector [0,1,0] v* cltn2_rep J)
   have "proj2_incident ?s (polar p)
      (vector [0,1,0] v* cltn2_rep J)
      (M *v (vector [1,0,1] v* cltn2_rep J)) = 0"
     and "proj2_incident ?s (polar q)
      (vector [0,1,0] v* cltn2_rep J)
      (M *v (vector [-1,0,1] v* cltn2_rep J)) = 0"
     by (simp_all add: proj2_incident_abs)
   with proj2_incident ?s (polar p) and proj2_incident ?s (polar q)
   have "(vector [0,1,0] v* cltn2_rep J)
      (M *v (vector [1,0,1] v* cltn2_rep J)) = 0"
     and "(vector [0,1,0] v* cltn2_rep J)
      (M *v (vector [-1,0,1] v* cltn2_rep J)) = 0"
     by simp_all
   hence "vector [0,1,0]  (?N *v vector [1,0,1]) = 0"
     and "vector [0,1,0]  (?N *v vector [-1,0,1]) = 0"
     by (simp_all add: dot_lmul_matrix matrix_vector_mul_assoc [symmetric])
   hence "?N$2$1 + ?N$2$3 = 0" and "-(?N$2$1) + ?N$2$3 = 0"
     unfolding inner_vec_def and matrix_vector_mult_def
     by (simp_all add: sum_3 vector_3)
   hence "?N$2$1 + ?N$2$3 = -(?N$2$1) + ?N$2$3" by simp
   hence "?N$2$1 = 0" by simp
   with ?N$2$1 + ?N$2$3 = 0 have "?N$2$3 = 0" by simp
   with ?N$2$2 + 2 * (?N$2$3) + ?N$3$3 = 0 and ?N$3$3 = -(?N$1$1)
   have "?N$2$2 = ?N$1$1" by simp
   with ?N$1$3 = 0 and ?N$2$1 = ?N$1$2 and ?N$1$3 = 0
     and ?N$2$1 = 0 and ?N$2$2 = ?N$1$1 and ?N$2$3 = 0
     and ?N$3$1 = ?N$1$3 and ?N$3$2 = ?N$2$3 and ?N$3$3 = -(?N$1$1)
   have "?N = (?N$1$1) *R M"
     unfolding M_def
     by (simp add: vec_eq_iff vector_3 forall_3)

   have "invertible (cltn2_rep J)" by (rule cltn2_rep_invertible)
   with M_invertible
   have "invertible ?N" by (simp add: invertible_mult transpose_invertible)
   hence "?N  0" by (auto simp add: zero_not_invertible)
   with ?N = (?N$1$1) *R M have "?N$1$1  0" by auto
   with ?N = (?N$1$1) *R M
   have "is_K2_isometry (cltn2_abs (cltn2_rep J))"
     by (simp add: J_M_J_transpose_K2_isometry)
   hence "is_K2_isometry J" by (simp add: cltn2_abs_rep)
   with apply_cltn2 east J = p
     and apply_cltn2 west J = q
     and apply_cltn2 north J = r
     and apply_cltn2 far_north J = ?s
   show " J. is_K2_isometry J
      apply_cltn2 east J = p
      apply_cltn2 west J = q
      apply_cltn2 north J = r
      apply_cltn2 far_north J = ?s"
     by auto
qed

lemma statement66_existence:
  assumes "a1  K2" and "a2  K2" and "p1  S" and "p2  S"
  shows " J. is_K2_isometry J  apply_cltn2 a1 J = a2  apply_cltn2 p1 J = p2"
proof -
  let ?a = "vector [a1,a2] :: proj2^2"
  from a1  K2 and a2  K2 have " i. ?a$i  K2" by (simp add: forall_2)

  let ?p = "vector [p1,p2] :: proj2^2"
  from p1  S and p2  S have " i. ?p$i  S" by (simp add: forall_2)

  let ?l = "χ i. proj2_line_through (?a$i) (?p$i)"
  have " i. proj2_incident (?a$i) (?l$i)"
    by (simp add: proj2_line_through_incident)
  hence "proj2_incident (?a$1) (?l$1)" and "proj2_incident (?a$2) (?l$2)"
    by fast+

  have " i. proj2_incident (?p$i) (?l$i)"
    by (simp add: proj2_line_through_incident)
  hence "proj2_incident (?p$1) (?l$1)" and "proj2_incident (?p$2) (?l$2)"
    by fast+

  let ?q = "χ i. ϵ qi. qi  ?p$i  qi  S  proj2_incident qi (?l$i)"
  have " i. ?q$i  ?p$i  ?q$i  S  proj2_incident (?q$i) (?l$i)"
  proof
    fix i
    from  i. ?a$i  K2 have "?a$i  K2" ..

    from  i. proj2_incident (?a$i) (?l$i)
    have "proj2_incident (?a$i) (?l$i)" ..
    with ?a$i  K2
    have " qi. qi  ?p$i  qi  S  proj2_incident qi (?l$i)"
      by (rule line_through_K2_intersect_S_again)
    with someI_ex [of "λ qi. qi  ?p$i  qi  S  proj2_incident qi (?l$i)"]
    show "?q$i  ?p$i  ?q$i  S  proj2_incident (?q$i) (?l$i)" by simp
  qed
  hence "?q$1  ?p$1" and "proj2_incident (?q$1) (?l$1)"
    and "proj2_incident (?q$2) (?l$2)"
    by fast+

  let ?r = "χ i. proj2_intersection (polar (?q$i)) (polar (?p$i))"
  let ?m = "χ i. proj2_line_through (?a$i) (?r$i)"
  have " i. proj2_incident (?a$i) (?m$i)"
    by (simp add: proj2_line_through_incident)
  hence "proj2_incident (?a$1) (?m$1)" and "proj2_incident (?a$2) (?m$2)"
    by fast+

  have " i. proj2_incident (?r$i) (?m$i)"
    by (simp add: proj2_line_through_incident)
  hence "proj2_incident (?r$1) (?m$1)" and "proj2_incident (?r$2) (?m$2)"
    by fast+

  let ?s = "χ i. ϵ si. si  ?r$i  si  S  proj2_incident si (?m$i)"
  have " i. ?s$i  ?r$i  ?s$i  S  proj2_incident (?s$i) (?m$i)"
  proof
    fix i
    from  i. ?a$i  K2 have "?a$i  K2" ..

    from  i. proj2_incident (?a$i) (?m$i)
    have "proj2_incident (?a$i) (?m$i)" ..
    with ?a$i  K2
    have " si. si  ?r$i  si  S  proj2_incident si (?m$i)"
      by (rule line_through_K2_intersect_S_again)
    with someI_ex [of "λ si. si  ?r$i  si  S  proj2_incident si (?m$i)"]
    show "?s$i  ?r$i  ?s$i  S  proj2_incident (?s$i) (?m$i)" by simp
  qed
  hence "?s$1  ?r$1" and "proj2_incident (?s$1) (?m$1)"
    and "proj2_incident (?s$2) (?m$2)"
    by fast+

  have " i .  u. proj2_incident u (?m$i)  ¬ (u = ?p$i  u = ?q$i)"
  proof standard+
    fix i :: 2
    fix u :: proj2
    assume "proj2_incident u (?m$i)"
    assume "u = ?p$i  u = ?q$i"

    from  i. ?p$i  S have "?p$i  S" ..

    from  i. ?q$i  ?p$i  ?q$i  S  proj2_incident (?q$i) (?l$i)
    have "?q$i  ?p$i" and "?q$i  S"
      by simp_all

    from ?p$i  S and ?q$i  S and u = ?p$i  u = ?q$i
    have "u  S" by auto
    hence "proj2_incident u (polar u)"
      by (simp add: incident_own_polar_in_S)

    have "proj2_incident (?r$i) (polar (?p$i))"
      and "proj2_incident (?r$i) (polar (?q$i))"
      by (simp_all add: proj2_intersection_incident)
    with u = ?p$i  u = ?q$i
    have "proj2_incident (?r$i) (polar u)" by auto

    from  i. proj2_incident (?r$i) (?m$i)
    have "proj2_incident (?r$i) (?m$i)" ..

    from  i. proj2_incident (?a$i) (?m$i)
    have "proj2_incident (?a$i) (?m$i)" ..

    from  i. ?a$i  K2 have "?a$i  K2" ..

    have "u  ?r$i"
    proof
      assume "u = ?r$i"
      with proj2_incident (?r$i) (polar (?p$i))
        and proj2_incident (?r$i) (polar (?q$i))
      have "proj2_incident u (polar (?p$i))"
        and "proj2_incident u (polar (?q$i))"
        by simp_all
      with u  S and ?p$i  S and ?q$i  S
      have "u = ?p$i" and "u = ?q$i"
        by (simp_all add: point_in_S_polar_is_tangent)
      with ?q$i  ?p$i show False by simp
    qed
    with proj2_incident (u) (polar u)
      and proj2_incident (?r$i) (polar u)
      and proj2_incident u (?m$i)
      and proj2_incident (?r$i) (?m$i)
      and proj2_incident_unique
    have "?m$i = polar u" by auto
    with proj2_incident (?a$i) (?m$i)
    have "proj2_incident (?a$i) (polar u)" by simp
    with u  S and ?a$i  K2 and tangent_not_through_K2
    show False by simp
  qed

  let ?H = "χ i. ϵ Hi. is_K2_isometry Hi
     apply_cltn2 east Hi = ?q$i
     apply_cltn2 west Hi = ?p$i
     apply_cltn2 north Hi = ?s$i
     apply_cltn2 far_north Hi = ?r$i"
  have " i. is_K2_isometry (?H$i)
     apply_cltn2 east (?H$i) = ?q$i
     apply_cltn2 west (?H$i) = ?p$i
     apply_cltn2 north (?H$i) = ?s$i
     apply_cltn2 far_north (?H$i) = ?r$i"
  proof
    fix i :: 2
    from  i. ?p$i  S have "?p$i  S" ..

    from  i. ?q$i  ?p$i  ?q$i  S  proj2_incident (?q$i) (?l$i)
    have "?q$i  ?p$i" and "?q$i  S"
      by simp_all

    from  i. ?s$i  ?r$i  ?s$i  S  proj2_incident (?s$i) (?m$i)
    have "?s$i  S" and "proj2_incident (?s$i) (?m$i)" by simp_all
    from proj2_incident (?s$i) (?m$i)
      and  i.  u.  proj2_incident u (?m$i)  ¬ (u = ?p$i  u = ?q$i)
    have "?s$i  {?q$i, ?p$i}" by fast
    with ?q$i  S and ?p$i  S and ?s$i  S and ?q$i  ?p$i
    have " Hi. is_K2_isometry Hi
       apply_cltn2 east Hi = ?q$i
       apply_cltn2 west Hi = ?p$i
       apply_cltn2 north Hi = ?s$i
       apply_cltn2 far_north Hi = ?r$i"
      by (simp add: statement65_special_case)
    with someI_ex [of "λ Hi. is_K2_isometry Hi
       apply_cltn2 east Hi = ?q$i
       apply_cltn2 west Hi = ?p$i
       apply_cltn2 north Hi = ?s$i
       apply_cltn2 far_north Hi = ?r$i"]
    show "is_K2_isometry (?H$i)
       apply_cltn2 east (?H$i) = ?q$i
       apply_cltn2 west (?H$i) = ?p$i
       apply_cltn2 north (?H$i) = ?s$i
       apply_cltn2 far_north (?H$i) = ?r$i"
      by simp
  qed
  hence "is_K2_isometry (?H$1)"
    and "apply_cltn2 east (?H$1) = ?q$1"
    and "apply_cltn2 west (?H$1) = ?p$1"
    and "apply_cltn2 north (?H$1) = ?s$1"
    and "apply_cltn2 far_north (?H$1) = ?r$1"
    and  "is_K2_isometry (?H$2)"
    and "apply_cltn2 east (?H$2) = ?q$2"
    and "apply_cltn2 west (?H$2) = ?p$2"
    and "apply_cltn2 north (?H$2) = ?s$2"
    and "apply_cltn2 far_north (?H$2) = ?r$2"
    by fast+

  let ?J = "cltn2_compose (cltn2_inverse (?H$1)) (?H$2)"
  from is_K2_isometry (?H$1) and is_K2_isometry (?H$2)
  have "is_K2_isometry ?J"
    by (simp only: cltn2_inverse_is_K2_isometry cltn2_compose_is_K2_isometry)

  from apply_cltn2 west (?H$1) = ?p$1
  have "apply_cltn2 p1 (cltn2_inverse (?H$1)) = west"
    by (simp add: cltn2.act_inv_iff [simplified])
  with apply_cltn2 west (?H$2) = ?p$2
  have "apply_cltn2 p1 ?J = p2"
    by (simp add: cltn2.act_act [simplified, symmetric])

  from apply_cltn2 east (?H$1) = ?q$1
  have "apply_cltn2 (?q$1) (cltn2_inverse (?H$1)) = east"
    by (simp add: cltn2.act_inv_iff [simplified])
  with apply_cltn2 east (?H$2) = ?q$2
  have "apply_cltn2 (?q$1) ?J = ?q$2"
    by (simp add: cltn2.act_act [simplified, symmetric])
  with ?q$1  ?p$1 and apply_cltn2 p1 ?J = p2
    and proj2_incident (?p$1) (?l$1)
    and proj2_incident (?q$1) (?l$1)
    and proj2_incident (?p$2) (?l$2)
    and proj2_incident (?q$2) (?l$2)
  have "apply_cltn2_line (?l$1) ?J = (?l$2)"
    by (simp add: apply_cltn2_line_unique)
  moreover from proj2_incident (?a$1) (?l$1)
  have "proj2_incident (apply_cltn2 (?a$1) ?J) (apply_cltn2_line (?l$1) ?J)"
    by simp
  ultimately have "proj2_incident (apply_cltn2 (?a$1) ?J) (?l$2)" by simp

  from apply_cltn2 north (?H$1) = ?s$1
  have "apply_cltn2 (?s$1) (cltn2_inverse (?H$1)) = north"
    by (simp add: cltn2.act_inv_iff [simplified])
  with apply_cltn2 north (?H$2) = ?s$2
  have "apply_cltn2 (?s$1) ?J = ?s$2"
    by (simp add: cltn2.act_act [simplified, symmetric])

  from apply_cltn2 far_north (?H$1) = ?r$1
  have "apply_cltn2 (?r$1) (cltn2_inverse (?H$1)) = far_north"
    by (simp add: cltn2.act_inv_iff [simplified])
  with apply_cltn2 far_north (?H$2) = ?r$2
  have "apply_cltn2 (?r$1) ?J = ?r$2"
    by (simp add: cltn2.act_act [simplified, symmetric])
  with ?s$1  ?r$1 and apply_cltn2 (?s$1) ?J = (?s$2)
    and proj2_incident (?r$1) (?m$1)
    and proj2_incident (?s$1) (?m$1)
    and proj2_incident (?r$2) (?m$2)
    and proj2_incident (?s$2) (?m$2)
  have "apply_cltn2_line (?m$1) ?J = (?m$2)"
    by (simp add: apply_cltn2_line_unique)
  moreover from proj2_incident (?a$1) (?m$1)
  have "proj2_incident (apply_cltn2 (?a$1) ?J) (apply_cltn2_line (?m$1) ?J)"
    by simp
  ultimately have "proj2_incident (apply_cltn2 (?a$1) ?J) (?m$2)" by simp

  from  i.  u. proj2_incident u (?m$i)  ¬ (u = ?p$i  u = ?q$i)
  have "¬ proj2_incident (?p$2) (?m$2)" by fast
  with proj2_incident (?p$2) (?l$2) have "?m$2  ?l$2" by auto
  with proj2_incident (?a$2) (?l$2)
    and proj2_incident (?a$2) (?m$2)
    and proj2_incident (apply_cltn2 (?a$1) ?J) (?l$2)
    and proj2_incident (apply_cltn2 (?a$1) ?J) (?m$2)
    and proj2_incident_unique
  have "apply_cltn2 a1 ?J = a2" by auto
  with is_K2_isometry ?J and apply_cltn2 p1 ?J = p2
  show " J. is_K2_isometry J  apply_cltn2 a1 J = a2  apply_cltn2 p1 J = p2"
    by auto
qed

lemma K2_isometry_swap:
  assumes "a  hyp2" and "b  hyp2"
  shows " J. is_K2_isometry J  apply_cltn2 a J = b  apply_cltn2 b J = a"
proof -
  from a  hyp2 and b  hyp2
  have "a  K2" and "b  K2" by simp_all

  let ?l = "proj2_line_through a b"
  have "proj2_incident a ?l" and "proj2_incident b ?l"
    by (rule proj2_line_through_incident)+
  from a  K2 and proj2_incident a ?l
    and line_through_K2_intersect_S_exactly_twice [of a ?l]
  obtain p and q where "p  q"
    and "p  S" and "q  S"
    and "proj2_incident p ?l" and "proj2_incident q ?l"
    and " rS. proj2_incident r ?l  r = p  r = q"
    by auto
  from a  K2 and b  K2 and p  S and q  S
    and statement66_existence [of a b p q]
  obtain J where "is_K2_isometry J" and "apply_cltn2 a J = b"
    and "apply_cltn2 p J = q"
    by auto
  from apply_cltn2 a J = b and apply_cltn2 p J = q
    and proj2_incident b ?l and proj2_incident q ?l
  have "proj2_incident (apply_cltn2 a J) ?l"
    and "proj2_incident (apply_cltn2 p J) ?l"
    by simp_all

  from a  K2 and p  S have "a  p"
    unfolding S_def and K2_def
    by auto
  with proj2_incident a ?l
    and proj2_incident p ?l
    and proj2_incident (apply_cltn2 a J) ?l
    and proj2_incident (apply_cltn2 p J) ?l
  have "apply_cltn2_line ?l J = ?l" by (simp add: apply_cltn2_line_unique)
  with proj2_incident q ?l and apply_cltn2_preserve_incident [of q J ?l]
  have "proj2_incident (apply_cltn2 q J) ?l" by simp

  from q  S and is_K2_isometry J
  have "apply_cltn2 q J  S" by (unfold is_K2_isometry_def) simp
  with proj2_incident (apply_cltn2 q J) ?l
    and  rS. proj2_incident r ?l  r = p  r = q
  have "apply_cltn2 q J = p  apply_cltn2 q J = q" by simp

  have "apply_cltn2 q J  q"
  proof
    assume "apply_cltn2 q J = q"
    with apply_cltn2 p J = q
    have "apply_cltn2 p J = apply_cltn2 q J" by simp
    hence "p = q" by (rule apply_cltn2_injective [of p J q])
    with p  q show False ..
  qed
  with apply_cltn2 q J = p  apply_cltn2 q J = q
  have "apply_cltn2 q J = p" by simp
  with p  q
    and apply_cltn2 p J = q
    and proj2_incident p ?l
    and proj2_incident q ?l
    and proj2_incident a ?l
    and statement55
  have "apply_cltn2 (apply_cltn2 a J) J = a" by simp
  with apply_cltn2 a J = b have "apply_cltn2 b J = a" by simp
  with is_K2_isometry J and apply_cltn2 a J = b
  show " J. is_K2_isometry J  apply_cltn2 a J = b  apply_cltn2 b J = a"
    by (simp add: exI [of _ J])
qed

theorem hyp2_axiom1: " a b. a b \<congruent>K b a"
proof standard+
  fix a b
  let ?a' = "Rep_hyp2 a"
  let ?b' = "Rep_hyp2 b"
  from Rep_hyp2 and K2_isometry_swap [of ?a' ?b']
  obtain J where "is_K2_isometry J" and "apply_cltn2 ?a' J = ?b'"
    and "apply_cltn2 ?b' J = ?a'"
    by auto

  from apply_cltn2 ?a' J = ?b' and apply_cltn2 ?b' J = ?a'
  have "hyp2_cltn2 a J = b" and "hyp2_cltn2 b J = a"
    unfolding hyp2_cltn2_def by (simp_all add: Rep_hyp2_inverse)
  with is_K2_isometry J
  show "a b \<congruent>K b a"
    by (unfold real_hyp2_C_def) (simp add: exI [of _ J])
qed

theorem hyp2_axiom2: " a b p q r s. a b \<congruent>K p q  a b \<congruent>K r s  p q \<congruent>K r s"
proof standard+
  fix a b p q r s
  assume "a b \<congruent>K p q  a b \<congruent>K r s"
  then obtain G and H where "is_K2_isometry G" and "is_K2_isometry H"
    and "hyp2_cltn2 a G = p" and "hyp2_cltn2 b G = q"
   and "hyp2_cltn2 a H = r" and "hyp2_cltn2 b H = s"
    by (unfold real_hyp2_C_def) auto
  let ?J = "cltn2_compose (cltn2_inverse G) H"
  from is_K2_isometry G have "is_K2_isometry (cltn2_inverse G)"
    by (rule cltn2_inverse_is_K2_isometry)
  with is_K2_isometry H
  have "is_K2_isometry ?J" by (simp only: cltn2_compose_is_K2_isometry)

  from is_K2_isometry G and hyp2_cltn2 a G = p and hyp2_cltn2 b G = q
    and K2_isometry.act_inv_iff
  have "hyp2_cltn2 p (cltn2_inverse G) = a"
    and "hyp2_cltn2 q (cltn2_inverse G) = b"
    by simp_all
  with hyp2_cltn2 a H = r and hyp2_cltn2 b H = s
    and is_K2_isometry (cltn2_inverse G) and is_K2_isometry H
    and K2_isometry.act_act [symmetric]
  have "hyp2_cltn2 p ?J = r" and "hyp2_cltn2 q ?J = s" by simp_all
  with is_K2_isometry ?J
  show "p q \<congruent>K r s"
    by (unfold real_hyp2_C_def) (simp add: exI [of _ ?J])
qed

theorem hyp2_axiom3: " a b c. a b \<congruent>K c c  a = b"
proof standard+
  fix a b c
  assume "a b \<congruent>K c c"
  then obtain J where "is_K2_isometry J"
    and "hyp2_cltn2 a J = c" and "hyp2_cltn2 b J = c"
    by (unfold real_hyp2_C_def) auto
  from hyp2_cltn2 a J = c and hyp2_cltn2 b J = c
  have "hyp2_cltn2 a J = hyp2_cltn2 b J" by simp

  from is_K2_isometry J
  have "apply_cltn2 (Rep_hyp2 a) J  hyp2"
    and "apply_cltn2 (Rep_hyp2 b) J  hyp2"
    by (rule apply_cltn2_Rep_hyp2)+
  with hyp2_cltn2 a J = hyp2_cltn2 b J
  have "apply_cltn2 (Rep_hyp2 a) J = apply_cltn2 (Rep_hyp2 b) J"
    by (unfold hyp2_cltn2_def) (simp add: Abs_hyp2_inject)
  hence "Rep_hyp2 a = Rep_hyp2 b" by (rule apply_cltn2_injective)
  thus "a = b" by (simp add: Rep_hyp2_inject)
qed

interpretation hyp2: tarski_first3 real_hyp2_C
  using hyp2_axiom1 and hyp2_axiom2 and hyp2_axiom3
  by unfold_locales

subsection ‹Some lemmas about betweenness›

lemma S_at_edge:
  assumes "p  S" and "q  hyp2  S" and "r  hyp2  S" and "proj2_Col p q r"
  shows "B (cart2_pt p) (cart2_pt q) (cart2_pt r)
   B (cart2_pt p) (cart2_pt r) (cart2_pt q)"
  (is "B ?cp ?cq ?cr  _")
proof -
  from p  S and q  hyp2  S and r  hyp2  S
  have "z_non_zero p" and "z_non_zero q" and "z_non_zero r"
    by (simp_all add: hyp2_S_z_non_zero)
  with proj2_Col p q r
  have "real_euclid.Col ?cp ?cq ?cr" by (simp add: proj2_Col_iff_euclid_cart2)

  with z_non_zero p and z_non_zero q and z_non_zero r
  have "proj2_pt ?cp = p" and "proj2_pt ?cq = q" and "proj2_pt ?cr = r"
    by (simp_all add: proj2_cart2)
  from proj2_pt ?cp = p and p  S
  have "norm ?cp = 1" by (simp add: norm_eq_1_iff_in_S)

  from proj2_pt ?cq = q and proj2_pt ?cr = r
    and q  hyp2  S and r  hyp2  S
  have "norm ?cq  1" and "norm ?cr  1"
    by (simp_all add: norm_le_1_iff_in_hyp2_S)

  show "B ?cp ?cq ?cr  B ?cp ?cr ?cq"
  proof cases
    assume "B ?cr ?cp ?cq"
    then obtain k where "k  0" and "k  1"
      and "?cp - ?cr = k *R (?cq - ?cr)"
      by (unfold real_euclid_B_def) auto
    from ?cp - ?cr = k *R (?cq - ?cr)
    have "?cp = k *R ?cq + (1 - k) *R ?cr" by (simp add: algebra_simps)
    with norm ?cp = 1 have "norm (k *R ?cq + (1 - k) *R ?cr) = 1" by simp
    with norm_triangle_ineq [of "k *R ?cq" "(1 - k) *R ?cr"]
    have "norm (k *R ?cq) + norm ((1 - k) *R ?cr)  1" by simp

    from k  0 and k  1
    have "norm (k *R ?cq) + norm ((1 - k) *R ?cr)
      = k * norm ?cq + (1 - k) * norm ?cr"
      by simp
    with norm (k *R ?cq) + norm ((1 - k) *R ?cr)  1
    have "k * norm ?cq + (1 - k) * norm ?cr  1" by simp

    from norm ?cq  1 and k  0 and mult_mono [of k k "norm ?cq" 1]
    have "k * norm ?cq  k" by simp

    from norm ?cr  1 and k  1
      and mult_mono [of "1 - k" "1 - k" "norm ?cr" 1]
    have "(1 - k) * norm ?cr  1 - k" by simp
    with k * norm ?cq  k
    have "k * norm ?cq + (1 - k) * norm ?cr  1" by simp
    with k * norm ?cq + (1 - k) * norm ?cr  1
    have "k * norm ?cq + (1 - k) * norm ?cr = 1" by simp
    with k * norm ?cq  k have "(1 - k) * norm ?cr  1 - k" by simp
    with (1 - k) * norm ?cr  1 - k have "(1 - k) * norm ?cr = 1 - k" by simp
    with k * norm ?cq + (1 - k) * norm ?cr = 1 have "k * norm ?cq = k" by simp

    have "?cp = ?cq  ?cq = ?cr  ?cr = ?cp"
    proof cases
      assume "k = 0  k = 1"
      with ?cp = k *R ?cq + (1 - k) *R ?cr
      show "?cp = ?cq  ?cq = ?cr  ?cr = ?cp" by auto
    next
      assume "¬ (k = 0  k = 1)"
      hence "k  0" and "k  1" by simp_all
      with k * norm ?cq = k and (1 - k) * norm ?cr = 1 - k
      have "norm ?cq = 1" and "norm ?cr = 1" by simp_all
      with proj2_pt ?cq = q and proj2_pt ?cr = r
      have "q  S" and "r  S" by (simp_all add: norm_eq_1_iff_in_S)
      with p  S have "{p,q,r}  S" by simp

      from proj2_Col p q r
      have "proj2_set_Col {p,q,r}" by (simp add: proj2_Col_iff_set_Col)
      with {p,q,r}  S have "card {p,q,r}  2" by (rule card_line_intersect_S)

      have "p = q  q = r  r = p"
      proof (rule ccontr)
        assume "¬ (p = q  q = r  r = p)"
        hence "p  q" and "q  r" and "r  p" by simp_all
        from q  r have "card {q,r} = 2" by simp
        with p  q and r  p have "card {p,q,r} = 3" by simp
        with card {p,q,r}  2 show False by simp
      qed
      thus "?cp = ?cq  ?cq = ?cr  ?cr = ?cp" by auto
    qed
    thus "B ?cp ?cq ?cr  B ?cp ?cr ?cq"
      by (auto simp add: real_euclid.th3_1 real_euclid.th3_2)
  next
    assume "¬ B ?cr ?cp ?cq"
    with real_euclid.Col ?cp ?cq ?cr
    show "B ?cp ?cq ?cr  B ?cp ?cr ?cq"
      unfolding real_euclid.Col_def
      by (auto simp add: real_euclid.th3_1 real_euclid.th3_2)
  qed
qed

lemma hyp2_in_middle:
  assumes "p  S" and "q  S" and "r  hyp2  S" and "proj2_Col p q r"
  and "p  q"
  shows "B (cart2_pt p) (cart2_pt r) (cart2_pt q)" (is "B ?cp ?cr ?cq")
proof (rule ccontr)
  assume "¬ B ?cp ?cr ?cq"
  hence "¬ B ?cq ?cr ?cp"
    by (auto simp add: real_euclid.th3_2 [of ?cq ?cr ?cp])

  from p  S and q  S and r  hyp2  S and proj2_Col p q r
  have "B ?cp ?cq ?cr  B ?cp ?cr ?cq" by (simp add: S_at_edge)
  with ¬ B ?cp ?cr ?cq have "B ?cp ?cq ?cr" by simp

  from proj2_Col p q r and proj2_Col_permute have "proj2_Col q p r" by fast
  with q  S and p  S and r  hyp2  S
  have "B ?cq ?cp ?cr  B ?cq ?cr ?cp" by (simp add: S_at_edge)
  with ¬ B ?cq ?cr ?cp have "B ?cq ?cp ?cr" by simp
  with B ?cp ?cq ?cr have "?cp = ?cq" by (rule real_euclid.th3_4)
  hence "proj2_pt ?cp = proj2_pt ?cq" by simp

  from p  S and q  S
  have "z_non_zero p" and "z_non_zero q" by (simp_all add: hyp2_S_z_non_zero)
  hence "proj2_pt ?cp = p" and "proj2_pt ?cq = q" by (simp_all add: proj2_cart2)
  with proj2_pt ?cp = proj2_pt ?cq have "p = q" by simp
  with p  q show False ..
qed

lemma hyp2_incident_in_middle:
  assumes "p  q" and "p  S" and "q  S" and "a  hyp2  S"
  and "proj2_incident p l" and "proj2_incident q l" and "proj2_incident a l"
  shows "B (cart2_pt p) (cart2_pt a) (cart2_pt q)"
proof -
  from proj2_incident p l and proj2_incident q l and proj2_incident a l
  have "proj2_Col p q a" by (rule proj2_incident_Col)
  from p  S and q  S and a  hyp2  S and this and p  q
  show "B (cart2_pt p) (cart2_pt a) (cart2_pt q)"
    by (rule hyp2_in_middle)
qed

lemma extend_to_S:
  assumes "p  hyp2  S" and "q  hyp2  S"
  shows " rS. B (cart2_pt p) (cart2_pt q) (cart2_pt r)"
  (is " rS. B ?cp ?cq (cart2_pt r)")
proof cases
  assume "q  S"

  have "B ?cp ?cq ?cq" by (rule real_euclid.th3_1)
  with q  S show " rS. B ?cp ?cq (cart2_pt r)" by auto
next
  assume "q  S"
  with q  hyp2  S have "q  K2" by simp

  let ?l = "proj2_line_through p q"
  have "proj2_incident p ?l" and "proj2_incident q ?l"
    by (rule proj2_line_through_incident)+
  from q  K2 and proj2_incident q ?l
    and line_through_K2_intersect_S_twice [of q ?l]
  obtain s and t where "s  t" and "s  S" and "t  S"
    and "proj2_incident s ?l" and "proj2_incident t ?l"
    by auto
  let ?cs = "cart2_pt s"
  let ?ct = "cart2_pt t"

  from proj2_incident s ?l
    and proj2_incident t ?l
    and proj2_incident p ?l
    and proj2_incident q ?l
  have "proj2_Col s p q" and "proj2_Col t p q" and "proj2_Col s t q"
    by (simp_all add: proj2_incident_Col)
  from proj2_Col s p q and proj2_Col t p q
    and s  S and t  S and p  hyp2  S and q  hyp2  S
  have "B ?cs ?cp ?cq  B ?cs ?cq ?cp" and "B ?ct ?cp ?cq  B ?ct ?cq ?cp"
    by (simp_all add: S_at_edge)
  with real_euclid.th3_2
  have "B ?cq ?cp ?cs  B ?cp ?cq ?cs" and "B ?cq ?cp ?ct  B ?cp ?cq ?ct"
    by fast+

  from s  S and t  S and q  hyp2  S and proj2_Col s t q and s  t
  have "B ?cs ?cq ?ct" by (rule hyp2_in_middle)
  hence "B ?ct ?cq ?cs" by (rule real_euclid.th3_2)

  have "B ?cp ?cq ?cs  B ?cp ?cq ?ct"
  proof (rule ccontr)
    assume "¬ (B ?cp ?cq ?cs  B ?cp ?cq ?ct)"
    hence "¬ B ?cp ?cq ?cs" and "¬ B ?cp ?cq ?ct" by simp_all
    with B ?cq ?cp ?cs  B ?cp ?cq ?cs
      and B ?cq ?cp ?ct  B ?cp ?cq ?ct
    have "B ?cq ?cp ?cs" and "B ?cq ?cp ?ct" by simp_all
    from ¬ B ?cp ?cq ?cs and B ?cq ?cp ?cs have "?cp  ?cq" by auto
    with B ?cq ?cp ?cs and B ?cq ?cp ?ct
    have "B ?cq ?cs ?ct  B ?cq ?ct ?cs"
      by (simp add: real_euclid_th5_1 [of ?cq ?cp ?cs ?ct])
    with B ?cs ?cq ?ct and B ?ct ?cq ?cs
    have "?cq = ?cs  ?cq = ?ct" by (auto simp add: real_euclid.th3_4)
    with q  hyp2  S and s  S and t  S
    have "q = s  q = t" by (auto simp add: hyp2_S_cart2_inj)
    with s  S and t  S have "q  S" by auto
    with q  S show False ..
  qed
  with s  S and t  S show " rS. B ?cp ?cq (cart2_pt r)" by auto
qed

definition endpoint_in_S :: "proj2  proj2  proj2" where
  "endpoint_in_S a b
   ϵ p. pS  B (cart2_pt a) (cart2_pt b) (cart2_pt p)"

lemma endpoint_in_S:
  assumes "a  hyp2  S" and "b  hyp2  S"
  shows "endpoint_in_S a b  S" (is "?p  S")
  and "B (cart2_pt a) (cart2_pt b) (cart2_pt (endpoint_in_S a b))"
  (is "B ?ca ?cb ?cp")
proof -
  from a  hyp2  S and b  hyp2  S and extend_to_S
  have " p. p  S  B ?ca ?cb (cart2_pt p)" by auto
  hence "?p  S  B ?ca ?cb ?cp"
    by (unfold endpoint_in_S_def) (rule someI_ex)
  thus "?p  S" and "B ?ca ?cb ?cp" by simp_all
qed

lemma endpoint_in_S_swap:
  assumes "a  b" and "a  hyp2  S" and "b  hyp2  S"
  shows "endpoint_in_S a b  endpoint_in_S b a" (is "?p  ?q")
proof
  let ?ca = "cart2_pt a"
  let ?cb = "cart2_pt b"
  let ?cp = "cart2_pt ?p"
  let ?cq = "cart2_pt ?q"
  from a  b and a  hyp2  S and b  hyp2  S
  have "B ?ca ?cb ?cp" and "B ?cb ?ca ?cq"
    by (simp_all add: endpoint_in_S)

  assume "?p = ?q"
  with B ?cb ?ca ?cq have "B ?cb ?ca ?cp" by simp
  with B ?ca ?cb ?cp have "?ca = ?cb" by (rule real_euclid.th3_4)
  with a  hyp2  S and b  hyp2  S have "a = b" by (rule hyp2_S_cart2_inj)
  with a  b show False ..
qed

lemma endpoint_in_S_incident:
  assumes "a  b" and "a  hyp2  S" and "b  hyp2  S"
  and "proj2_incident a l" and "proj2_incident b l"
  shows "proj2_incident (endpoint_in_S a b) l" (is "proj2_incident ?p l")
proof -
  from a  hyp2  S and b  hyp2  S
  have "?p  S" and "B (cart2_pt a) (cart2_pt b) (cart2_pt ?p)"
    (is "B ?ca ?cb ?cp")
    by (rule endpoint_in_S)+

  from a  hyp2  S and b  hyp2  S and ?p  S
  have "z_non_zero a" and "z_non_zero b" and "z_non_zero ?p"
    by (simp_all add: hyp2_S_z_non_zero)

  from B ?ca ?cb ?cp
  have "real_euclid.Col ?ca ?cb ?cp" unfolding real_euclid.Col_def ..
  with z_non_zero a and z_non_zero b and z_non_zero ?p and a  b
    and proj2_incident a l and proj2_incident b l
  show "proj2_incident ?p l" by (rule euclid_Col_cart2_incident)
qed

lemma endpoints_in_S_incident_unique:
  assumes "a  b" and "a  hyp2  S" and "b  hyp2  S" and "p  S"
  and "proj2_incident a l" and "proj2_incident b l" and "proj2_incident p l"
  shows "p = endpoint_in_S a b  p = endpoint_in_S b a"
  (is "p = ?q  p = ?r")
proof -
  from a  b and a  hyp2  S and b  hyp2  S
  have "?q  ?r" by (rule endpoint_in_S_swap)

  from a  hyp2  S and b  hyp2  S
  have "?q  S" and "?r  S" by (simp_all add: endpoint_in_S)

  from a  b and a  hyp2  S and b  hyp2  S
    and proj2_incident a l and proj2_incident b l
  have "proj2_incident ?q l" and "proj2_incident ?r l"
    by (simp_all add: endpoint_in_S_incident)
  with ?q  ?r and ?q  S and ?r  S and p  S and proj2_incident p l
  show "p = ?q  p = ?r" by (simp add: line_S_two_intersections_only)
qed

lemma endpoint_in_S_unique:
  assumes "a  b" and  "a  hyp2  S" and "b  hyp2  S" and "p  S"
  and "B (cart2_pt a) (cart2_pt b) (cart2_pt p)" (is "B ?ca ?cb ?cp")
  shows "p = endpoint_in_S a b" (is "p = ?q")
proof (rule ccontr)
  from a  hyp2  S and b  hyp2  S and p  S
  have "z_non_zero a" and "z_non_zero b" and "z_non_zero p"
    by (simp_all add: hyp2_S_z_non_zero)
  with B ?ca ?cb ?cp and euclid_B_cart2_common_line [of a b p]
  obtain l where
    "proj2_incident a l" and "proj2_incident b l" and "proj2_incident p l"
    by auto
  with a  b and a  hyp2  S and b  hyp2  S and p  S
  have "p = ?q  p = endpoint_in_S b a" (is "p = ?q  p = ?r")
    by (rule endpoints_in_S_incident_unique)

  assume "p  ?q"
  with p = ?q  p = ?r have "p = ?r" by simp
  with b  hyp2  S and a  hyp2  S
  have "B ?cb ?ca ?cp" by (simp add: endpoint_in_S)
  with B ?ca ?cb ?cp have "?ca = ?cb" by (rule real_euclid.th3_4)
  with a  hyp2  S and b  hyp2  S have "a = b" by (rule hyp2_S_cart2_inj)
  with a  b show False ..
qed

lemma between_hyp2_S:
  assumes "p  hyp2  S" and "r  hyp2  S" and "k  0" and "k  1"
  shows "proj2_pt (k *R (cart2_pt r) + (1 - k) *R (cart2_pt p))  hyp2  S"
  (is "proj2_pt ?cq  _")
proof -
  let ?cp = "cart2_pt p"
  let ?cr = "cart2_pt r"
  let ?q = "proj2_pt ?cq"
  from p  hyp2  S and r  hyp2  S
  have "z_non_zero p" and "z_non_zero r" by (simp_all add: hyp2_S_z_non_zero)
  hence "proj2_pt ?cp = p" and "proj2_pt ?cr = r" by (simp_all add: proj2_cart2)
  with p  hyp2  S and r  hyp2  S
  have "norm ?cp  1" and "norm ?cr  1"
    by (simp_all add: norm_le_1_iff_in_hyp2_S)

  from k  0 and k  1
    and norm_triangle_ineq [of "k *R ?cr" "(1 - k) *R ?cp"]
  have "norm ?cq  k * norm ?cr + (1 - k) * norm ?cp" by simp

  from k  0 and norm ?cr  1 and mult_mono [of k k "norm ?cr" 1]
  have "k * norm ?cr  k" by simp

  from k  1 and norm ?cp  1
    and mult_mono [of "1 - k" "1 - k" "norm ?cp" 1]
  have "(1 - k) * norm ?cp  1 - k" by simp
  with norm ?cq  k * norm ?cr + (1 - k) * norm ?cp and k * norm ?cr  k
  have "norm ?cq  1" by simp
  thus "?q  hyp2  S" by (simp add: norm_le_1_iff_in_hyp2_S)
qed

subsection ‹The Klein--Beltrami model satisfies axiom 4›

definition expansion_factor :: "proj2  cltn2  real" where
  "expansion_factor p J  (cart2_append1 p v* cltn2_rep J)$3"

lemma expansion_factor:
  assumes "p  hyp2  S" and "is_K2_isometry J"
  shows "expansion_factor p J  0"
  and "cart2_append1 p v* cltn2_rep J
  = expansion_factor p J *R cart2_append1 (apply_cltn2 p J)"
proof -
  from p  hyp2  S and is_K2_isometry J
  have "z_non_zero (apply_cltn2 p J)" by (rule is_K2_isometry_z_non_zero)

  from p  hyp2  S and is_K2_isometry J
  and cart2_append1_apply_cltn2
  obtain k where "k  0"
    and "cart2_append1 p v* cltn2_rep J = k *R cart2_append1 (apply_cltn2 p J)"
    by auto
  from cart2_append1 p v* cltn2_rep J = k *R cart2_append1 (apply_cltn2 p J)
    and z_non_zero (apply_cltn2 p J)
  have "expansion_factor p J = k"
    by (unfold expansion_factor_def) (simp add: cart2_append1_z)
  with k  0
    and cart2_append1 p v* cltn2_rep J = k *R cart2_append1 (apply_cltn2 p J)
  show "expansion_factor p J  0"
    and "cart2_append1 p v* cltn2_rep J
    = expansion_factor p J *R cart2_append1 (apply_cltn2 p J)"
    by simp_all
qed

lemma expansion_factor_linear_apply_cltn2:
  assumes "p  hyp2  S" and "q  hyp2  S" and "r  hyp2  S"
  and "is_K2_isometry J"
  and "cart2_pt r = k *R cart2_pt p + (1 - k) *R cart2_pt q"
  shows "expansion_factor r J *R cart2_append1 (apply_cltn2 r J)
  = (k * expansion_factor p J) *R cart2_append1 (apply_cltn2 p J)
  + ((1 - k) * expansion_factor q J) *R cart2_append1 (apply_cltn2 q J)"
  (is "?er *R _ = (k * ?ep) *R _ + ((1 - k) * ?eq) *R _")
proof -
  let ?cp = "cart2_pt p"
  let ?cq = "cart2_pt q"
  let ?cr = "cart2_pt r"
  let ?cp1 = "cart2_append1 p"
  let ?cq1 = "cart2_append1 q"
  let ?cr1 = "cart2_append1 r"
  let ?repJ = "cltn2_rep J"
  from p  hyp2  S and q  hyp2  S and r  hyp2  S
  have "z_non_zero p" and "z_non_zero q" and "z_non_zero r"
    by (simp_all add: hyp2_S_z_non_zero)

  from ?cr = k *R ?cp + (1 - k) *R ?cq
  have "vector2_append1 ?cr
    = k *R vector2_append1 ?cp + (1 - k) *R vector2_append1 ?cq"
    by (unfold vector2_append1_def vector_def) (simp add: vec_eq_iff)
  with z_non_zero p and z_non_zero q and z_non_zero r
  have "?cr1 = k *R ?cp1 + (1 - k) *R ?cq1" by (simp add: cart2_append1)
  hence "?cr1 v* ?repJ = k *R (?cp1 v* ?repJ) + (1 - k) *R (?cq1 v* ?repJ)"
    by (simp add: vector_matrix_left_distrib scaleR_vector_matrix_assoc [symmetric])
  with p  hyp2  S and q  hyp2  S and r  hyp2  S
    and is_K2_isometry J
  show "?er *R cart2_append1 (apply_cltn2 r J)
    = (k * ?ep) *R cart2_append1 (apply_cltn2 p J)
    + ((1 - k) * ?eq) *R cart2_append1 (apply_cltn2 q J)"
    by (simp add: expansion_factor)
qed

lemma expansion_factor_linear:
  assumes "p  hyp2  S" and "q  hyp2  S" and "r  hyp2  S"
  and "is_K2_isometry J"
  and "cart2_pt r = k *R cart2_pt p + (1 - k) *R cart2_pt q"
  shows "expansion_factor r J
  = k * expansion_factor p J + (1 - k) * expansion_factor q J"
  (is "?er = k * ?ep + (1 - k) * ?eq")
proof -
  from p  hyp2  S and q  hyp2  S and r  hyp2  S
    and is_K2_isometry J
  have "z_non_zero (apply_cltn2 p J)"
    and "z_non_zero (apply_cltn2 q J)"
    and "z_non_zero (apply_cltn2 r J)"
    by (simp_all add: is_K2_isometry_z_non_zero)

  from p  hyp2  S and q  hyp2  S and r  hyp2  S
    and is_K2_isometry J
    and cart2_pt r = k *R cart2_pt p + (1 - k) *R cart2_pt q
  have "?er *R cart2_append1 (apply_cltn2 r J)
    = (k * ?ep) *R cart2_append1 (apply_cltn2 p J)
    + ((1 - k) * ?eq) *R cart2_append1 (apply_cltn2 q J)"
    by (rule expansion_factor_linear_apply_cltn2)
  hence "(?er *R cart2_append1 (apply_cltn2 r J))$3
    = ((k * ?ep) *R cart2_append1 (apply_cltn2 p J)
    + ((1 - k) * ?eq) *R cart2_append1 (apply_cltn2 q J))$3"
    by simp
  with z_non_zero (apply_cltn2 p J)
    and z_non_zero (apply_cltn2 q J)
    and z_non_zero (apply_cltn2 r J)
  show "?er = k * ?ep + (1 - k) * ?eq" by (simp add: cart2_append1_z)
qed

lemma expansion_factor_sgn_invariant:
  assumes "p  hyp2  S" and "q  hyp2  S" and "is_K2_isometry J"
  shows "sgn (expansion_factor p J) = sgn (expansion_factor q J)"
  (is "sgn ?ep = sgn ?eq")
proof (rule ccontr)
  assume "sgn ?ep  sgn ?eq"

  from p  hyp2  S and q  hyp2  S and is_K2_isometry J
  have "?ep  0" and "?eq  0" by (simp_all add: expansion_factor)
  hence "sgn ?ep  {-1,1}" and "sgn ?eq  {-1,1}"
    by (simp_all add: sgn_real_def)
  with sgn ?ep  sgn ?eq have "sgn ?ep = - sgn ?eq" by auto
  hence "sgn ?ep = sgn (-?eq)" by (subst sgn_minus)
  with sgn_plus [of ?ep "-?eq"]
  have "sgn (?ep - ?eq) = sgn ?ep" by (simp add: algebra_simps)
  with sgn ?ep  {-1,1} have "?ep - ?eq  0" by (auto simp add: sgn_real_def)

  let ?k = "-?eq / (?ep - ?eq)"
  from sgn (?ep - ?eq) = sgn ?ep and sgn ?ep = sgn (-?eq)
  have "sgn (?ep - ?eq) = sgn (-?eq)" by simp
  with ?ep - ?eq  0 and sgn_div [of "?ep - ?eq" "-?eq"]
  have "?k > 0" by simp

  from ?ep - ?eq  0
  have "1 - ?k = ?ep / (?ep - ?eq)" by (simp add: field_simps)
  with sgn (?ep - ?eq) = sgn ?ep and ?ep - ?eq  0
  have "1 - ?k > 0" by (simp add: sgn_div)
  hence "?k < 1" by simp

  let ?cp = "cart2_pt p"
  let ?cq = "cart2_pt q"
  let ?cr = "?k *R ?cp + (1 - ?k) *R ?cq"
  let ?r = "proj2_pt ?cr"
  let ?er = "expansion_factor ?r J"
  have "cart2_pt ?r = ?cr" by (rule cart2_proj2)

  from p  hyp2  S and q  hyp2  S and ?k > 0 and ?k < 1
    and between_hyp2_S [of q p ?k]
  have "?r  hyp2  S" by simp
  with p  hyp2  S and q  hyp2  S and is_K2_isometry J
    and cart2_pt ?r = ?cr
    and expansion_factor_linear [of p q ?r J ?k]
  have "?er = ?k * ?ep + (1 - ?k) * ?eq" by simp
  with ?ep - ?eq  0 have "?er = 0" by (simp add: field_simps)
  with ?r  hyp2  S and is_K2_isometry J
  show False by (simp add: expansion_factor)
qed

lemma statement_63:
  assumes "p  hyp2  S" and "q  hyp2  S" and "r  hyp2  S"
  and "is_K2_isometry J" and "B (cart2_pt p) (cart2_pt q) (cart2_pt r)"
  shows "B
  (cart2_pt (apply_cltn2 p J))
  (cart2_pt (apply_cltn2 q J))
  (cart2_pt (apply_cltn2 r J))"
proof -
  let ?cp = "cart2_pt p"
  let ?cq = "cart2_pt q"
  let ?cr = "cart2_pt r"
  let ?ep = "expansion_factor p J"
  let ?eq = "expansion_factor q J"
  let ?er = "expansion_factor r J"
  from q  hyp2  S and is_K2_isometry J
  have "?eq  0" by (rule expansion_factor)

  from p  hyp2  S and q  hyp2  S and r  hyp2  S
    and is_K2_isometry J and expansion_factor_sgn_invariant
  have "sgn ?ep = sgn ?eq" and "sgn ?er = sgn ?eq" by fast+
  with ?eq  0
  have "?ep / ?eq > 0" and "?er / ?eq > 0" by (simp_all add: sgn_div)

  from B ?cp ?cq ?cr
  obtain k where "k  0" and "k  1" and "?cq = k *R ?cr + (1 - k) *R ?cp"
    by (unfold real_euclid_B_def) (auto simp add: algebra_simps)

  let ?c = "k * ?er / ?eq"
  from k  0 and ?er / ?eq > 0 and mult_nonneg_nonneg [of k "?er / ?eq"]
  have "?c  0" by simp

  from r  hyp2  S and p  hyp2  S and q  hyp2  S
    and is_K2_isometry J and ?cq = k *R ?cr + (1 - k) *R ?cp
  have "?eq = k * ?er + (1 - k) * ?ep" by (rule expansion_factor_linear)
  with ?eq  0 have "1 - ?c = (1 - k) * ?ep / ?eq" by (simp add: field_simps)
  with k  1 and ?ep / ?eq > 0
    and mult_nonneg_nonneg [of "1 - k" "?ep / ?eq"]
  have "?c  1" by simp

  let ?pJ = "apply_cltn2 p J"
  let ?qJ = "apply_cltn2 q J"
  let ?rJ = "apply_cltn2 r J"
  let ?cpJ = "cart2_pt ?pJ"
  let ?cqJ = "cart2_pt ?qJ"
  let ?crJ = "cart2_pt ?rJ"
  let ?cpJ1 = "cart2_append1 ?pJ"
  let ?cqJ1 = "cart2_append1 ?qJ"
  let ?crJ1 = "cart2_append1 ?rJ"
  from p  hyp2  S and q  hyp2  S and r  hyp2  S
    and is_K2_isometry J
  have "z_non_zero ?pJ" and "z_non_zero ?qJ" and "z_non_zero ?rJ"
    by (simp_all add: is_K2_isometry_z_non_zero)

  from r  hyp2  S and p  hyp2  S and q  hyp2  S
    and is_K2_isometry J and ?cq = k *R ?cr + (1 - k) *R ?cp
  have "?eq *R ?cqJ1 = (k * ?er) *R ?crJ1 + ((1 - k) * ?ep) *R ?cpJ1"
    by (rule expansion_factor_linear_apply_cltn2)
  hence "(1 / ?eq) *R (?eq *R ?cqJ1)
    = (1 / ?eq) *R ((k * ?er) *R ?crJ1 + ((1 - k) * ?ep) *R ?cpJ1)" by simp
  with 1 - ?c = (1 - k) * ?ep / ?eq and ?eq  0
  have "?cqJ1 = ?c *R ?crJ1 + (1 - ?c) *R ?cpJ1"
    by (simp add: scaleR_right_distrib)
  with z_non_zero ?pJ and z_non_zero ?qJ and z_non_zero ?rJ
  have "vector2_append1 ?cqJ
    = ?c *R vector2_append1 ?crJ + (1 - ?c) *R vector2_append1 ?cpJ"
    by (simp add: cart2_append1)
  hence "?cqJ = ?c *R ?crJ + (1 - ?c) *R ?cpJ"
    unfolding vector2_append1_def and vector_def
    by (simp add: vec_eq_iff forall_2 forall_3)
  with ?c  0 and ?c  1
  show "B ?cpJ ?cqJ ?crJ"
    by (unfold real_euclid_B_def) (simp add: algebra_simps exI [of _ ?c])
qed

theorem hyp2_axiom4: " q a b c.  x. BK q a x  a x \<congruent>K b c"
proof (rule allI)+
  fix q a b c :: hyp2
  let ?pq = "Rep_hyp2 q"
  let ?pa = "Rep_hyp2 a"
  let ?pb = "Rep_hyp2 b"
  let ?pc = "Rep_hyp2 c"
  have "?pq  hyp2" and "?pa  hyp2" and "?pb  hyp2" and "?pc  hyp2"
    by (rule Rep_hyp2)+
  let ?cq = "cart2_pt ?pq"
  let ?ca = "cart2_pt ?pa"
  let ?cb = "cart2_pt ?pb"
  let ?cc = "cart2_pt ?pc"
  let ?pp = "ϵ p. p  S  B ?cb ?cc (cart2_pt p)"
  let ?cp = "cart2_pt ?pp"
  from ?pb  hyp2 and ?pc  hyp2 and extend_to_S [of ?pb ?pc]
    and someI_ex [of "λ p. p  S  B ?cb ?cc (cart2_pt p)"]
  have "?pp  S" and "B ?cb ?cc ?cp" by auto

  let ?pr = "ϵ r. r  S  B ?cq ?ca (cart2_pt r)"
  let ?cr = "cart2_pt ?pr"
  from ?pq  hyp2 and ?pa  hyp2 and extend_to_S [of ?pq ?pa]
    and someI_ex [of "λ r. r  S  B ?cq ?ca (cart2_pt r)"]
  have "?pr  S" and "B ?cq ?ca ?cr" by auto

  from ?pb  hyp2 and ?pa  hyp2 and ?pp  S and ?pr  S
    and statement66_existence [of ?pb ?pa ?pp ?pr]
  obtain J where "is_K2_isometry J"
    and "apply_cltn2 ?pb J = ?pa" and "apply_cltn2 ?pp J = ?pr"
    by auto
  let ?px = "apply_cltn2 ?pc J"
  let ?cx = "cart2_pt ?px"
  let ?x = "Abs_hyp2 ?px"
  from is_K2_isometry J and ?pc  hyp2
  have "?px  hyp2" by (rule statement60_one_way)
  hence "Rep_hyp2 ?x = ?px" by (rule Abs_hyp2_inverse)

  from ?pb  hyp2 and ?pc  hyp2 and ?pp  S and is_K2_isometry J
    and B ?cb ?cc ?cp and statement_63
  have "B (cart2_pt (apply_cltn2 ?pb J)) ?cx (cart2_pt (apply_cltn2 ?pp J))"
    by simp
  with apply_cltn2 ?pb J = ?pa and apply_cltn2 ?pp J = ?pr
  have "B ?ca ?cx ?cr" by simp
  with B ?cq ?ca ?cr have "B ?cq ?ca ?cx" by (rule real_euclid.th3_5_1)
  with Rep_hyp2 ?x = ?px
  have "BK q a ?x"
    unfolding real_hyp2_B_def and hyp2_rep_def
    by simp

  have "Abs_hyp2 ?pa = a" by (rule Rep_hyp2_inverse)
  with apply_cltn2 ?pb J = ?pa
  have "hyp2_cltn2 b J = a" by (unfold hyp2_cltn2_def) simp

  have "hyp2_cltn2 c J = ?x" unfolding hyp2_cltn2_def ..
  with is_K2_isometry J and hyp2_cltn2 b J = a
  have "b c \<congruent>K a ?x"
    by (unfold real_hyp2_C_def) (simp add: exI [of _ J])
  hence "a ?x \<congruent>K b c" by (rule hyp2.th2_2)
  with BK q a ?x
  show " x. BK q a x  a x \<congruent>K b c" by (simp add: exI [of _ ?x])
qed

subsection ‹More betweenness theorems›

lemma hyp2_S_points_fix_line:
  assumes "a  hyp2" and "p  S" and "is_K2_isometry J"
  and "apply_cltn2 a J = a" (is "?aJ = a")
  and "apply_cltn2 p J = p" (is "?pJ = p")
  and "proj2_incident a l" and "proj2_incident p l" and "proj2_incident b l"
  shows "apply_cltn2 b J = b" (is "?bJ = b")
proof -
  let ?lJ = "apply_cltn2_line l J"
  from proj2_incident a l and proj2_incident p l
  have "proj2_incident ?aJ ?lJ" and "proj2_incident ?pJ ?lJ" by simp_all
  with ?aJ = a and ?pJ = p
  have "proj2_incident a ?lJ" and "proj2_incident p ?lJ" by simp_all

  from a  hyp2 proj2_incident a l and line_through_K2_intersect_S_again [of a l]
  obtain q where "q  p" and "q  S" and "proj2_incident q l" by auto
  let ?qJ = "apply_cltn2 q J"

  from a  hyp2 and p  S and q  S
  have "a  p" and "a  q" by (simp_all add: hyp2_S_not_equal)

  from a  p and proj2_incident a l and proj2_incident p l
    and proj2_incident a ?lJ and proj2_incident p ?lJ
    and proj2_incident_unique
  have "?lJ = l" by auto

  from proj2_incident q l have "proj2_incident ?qJ ?lJ" by simp
  with ?lJ = l have "proj2_incident ?qJ l" by simp

  from q  S and is_K2_isometry J
  have "?qJ  S" by (unfold is_K2_isometry_def) simp
  with q  p and p  S and q  S and proj2_incident p l
    and proj2_incident q l and proj2_incident ?qJ l
    and line_S_two_intersections_only
  have "?qJ = p  ?qJ = q" by simp

  have "?qJ = q"
  proof (rule ccontr)
    assume "?qJ  q"
    with ?qJ = p  ?qJ = q have "?qJ = p" by simp
    with ?pJ = p have "?qJ = ?pJ" by simp
    with apply_cltn2_injective have "q = p" by fast
    with q  p show False ..
  qed
  with q  p and a  p and a  q and proj2_incident p l
    and proj2_incident q l and proj2_incident a l
    and ?pJ = p and ?aJ = a and proj2_incident b l
    and cltn2_three_point_line [of p q a l J b]
  show "?bJ = b" by simp
qed

lemma K2_isometry_endpoint_in_S:
  assumes "a  b" and "a  hyp2  S" and "b  hyp2  S" and "is_K2_isometry J"
  shows "apply_cltn2 (endpoint_in_S a b) J
  = endpoint_in_S (apply_cltn2 a J) (apply_cltn2 b J)"
  (is "?pJ = endpoint_in_S ?aJ ?bJ")
proof -
  let ?p = "endpoint_in_S a b"

  from a  b and apply_cltn2_injective have "?aJ  ?bJ" by fast

  from a  hyp2  S and b  hyp2  S and is_K2_isometry J
    and is_K2_isometry_hyp2_S
  have "?aJ  hyp2  S" and "?bJ  hyp2  S" by simp_all

  let ?ca = "cart2_pt a"
  let ?cb = "cart2_pt b"
  let ?cp = "cart2_pt ?p"
  from a  hyp2  S and b  hyp2  S
  have "?p  S" and "B ?ca ?cb ?cp" by (rule endpoint_in_S)+

  from ?p  S and is_K2_isometry J
  have "?pJ  S" by (unfold is_K2_isometry_def) simp

  let ?caJ = "cart2_pt ?aJ"
  let ?cbJ = "cart2_pt ?bJ"
  let ?cpJ = "cart2_pt ?pJ"
  from a  hyp2  S and b  hyp2  S and ?p  S and is_K2_isometry J
    and B ?ca ?cb ?cp and statement_63
  have "B ?caJ ?cbJ ?cpJ" by simp
  with ?aJ  ?bJ and ?aJ  hyp2  S and ?bJ  hyp2  S and ?pJ  S
  show "?pJ = endpoint_in_S ?aJ ?bJ" by (rule endpoint_in_S_unique)
qed

lemma between_endpoint_in_S:
  assumes "a  b" and "b  c"
  and "a  hyp2  S" and "b  hyp2  S" and "c  hyp2  S"
  and "B (cart2_pt a) (cart2_pt b) (cart2_pt c)" (is "B ?ca ?cb ?cc")
  shows "endpoint_in_S a b = endpoint_in_S b c" (is "?p = ?q")
proof -
  from b  c and b  hyp2  S and c  hyp2  S and hyp2_S_cart2_inj
  have "?cb  ?cc" by auto

  let ?cq = "cart2_pt ?q"
  from b  hyp2  S and c  hyp2  S
  have "?q  S" and "B ?cb ?cc ?cq" by (rule endpoint_in_S)+

  from ?cb  ?cc and B ?ca ?cb ?cc and B ?cb ?cc ?cq
  have "B ?ca ?cb ?cq" by (rule real_euclid.th3_7_2)
  with a  b and a  hyp2  S and b  hyp2  S and ?q  S
  have "?q = ?p" by (rule endpoint_in_S_unique)
  thus "?p = ?q" ..
qed

lemma hyp2_extend_segment_unique:
  assumes "a  b" and "BK a b c" and "BK a b d" and "b c \<congruent>K b d"
  shows "c = d"
proof cases
  assume "b = c"
  with b c \<congruent>K b d show "c = d" by (simp add: hyp2.A3_reversed)
next
  assume "b  c"

  have "b  d"
  proof (rule ccontr)
    assume "¬ b  d"
    hence "b = d" by simp
    with b c \<congruent>K b d have "b c \<congruent>K b b" by simp
    hence "b = c" by (rule hyp2.A3')
    with b  c show False ..
  qed
  with a  b and b  c
  have "Rep_hyp2 a  Rep_hyp2 b" (is "?pa  ?pb")
    and "Rep_hyp2 b  Rep_hyp2 c" (is "?pb  ?pc")
    and "Rep_hyp2 b  Rep_hyp2 d" (is "?pb  ?pd")
    by (simp_all add: Rep_hyp2_inject)

  have "?pa  hyp2" and "?pb  hyp2" and "?pc  hyp2" and "?pd  hyp2"
    by (rule Rep_hyp2)+

  let ?pp = "endpoint_in_S ?pb ?pc"
  let ?ca = "cart2_pt ?pa"
  let ?cb = "cart2_pt ?pb"
  let ?cc = "cart2_pt ?pc"
  let ?cd = "cart2_pt ?pd"
  let ?cp = "cart2_pt ?pp"
  from ?pb  hyp2 and ?pc  hyp2
  have "?pp  S" and "B ?cb ?cc ?cp" by (simp_all add: endpoint_in_S)

  from b c \<congruent>K b d
  obtain J where "is_K2_isometry J"
    and "hyp2_cltn2 b J = b" and "hyp2_cltn2 c J = d"
    by (unfold real_hyp2_C_def) auto

  from hyp2_cltn2 b J = b and hyp2_cltn2 c J = d
  have "Rep_hyp2 (hyp2_cltn2 b J) = ?pb"
    and "Rep_hyp2 (hyp2_cltn2 c J) = ?pd"
    by simp_all
  with is_K2_isometry J
  have "apply_cltn2 ?pb J = ?pb" and "apply_cltn2 ?pc J = ?pd"
    by (simp_all add: Rep_hyp2_cltn2)

  from BK a b c and BK a b d
  have "B ?ca ?cb ?cc" and "B ?ca ?cb ?cd"
    unfolding real_hyp2_B_def and hyp2_rep_def .

  from ?pb  ?pc and ?pb  hyp2 and ?pc  hyp2 and is_K2_isometry J
  have "apply_cltn2 ?pp J
    = endpoint_in_S (apply_cltn2 ?pb J) (apply_cltn2 ?pc J)"
    by (simp add: K2_isometry_endpoint_in_S)
  also from apply_cltn2 ?pb J = ?pb and apply_cltn2 ?pc J = ?pd
  have " = endpoint_in_S ?pb ?pd" by simp
  also from ?pa  ?pb and ?pb  ?pd
    and ?pa  hyp2 and ?pb  hyp2 and ?pd  hyp2 and B ?ca ?cb ?cd
  have " = endpoint_in_S ?pa ?pb" by (simp add: between_endpoint_in_S)
  also from ?pa  ?pb and ?pb  ?pc
    and ?pa  hyp2 and ?pb  hyp2 and ?pc  hyp2 and B ?ca ?cb ?cc
  have " = endpoint_in_S ?pb ?pc" by (simp add: between_endpoint_in_S)
  finally have "apply_cltn2 ?pp J = ?pp" .

  from ?pb  hyp2 and ?pc  hyp2 and ?pp  S
  have "z_non_zero ?pb" and "z_non_zero ?pc" and "z_non_zero ?pp"
    by (simp_all add: hyp2_S_z_non_zero)
  with B ?cb ?cc ?cp and euclid_B_cart2_common_line [of ?pb ?pc ?pp]
  obtain l where "proj2_incident ?pb l" and "proj2_incident ?pp l"
    and "proj2_incident ?pc l"
    by auto
  with ?pb  hyp2 and ?pp  S and is_K2_isometry J
    and apply_cltn2 ?pb J = ?pb and apply_cltn2 ?pp J = ?pp
  have "apply_cltn2 ?pc J = ?pc" by (rule hyp2_S_points_fix_line)
  with apply_cltn2 ?pc J = ?pd have "?pc = ?pd" by simp
  thus "c = d" by (subst Rep_hyp2_inject [symmetric])
qed

lemma line_S_match_intersections:
  assumes "p  q" and "r  s" and "p  S" and "q  S" and "r  S" and "s  S"
  and "proj2_set_Col {p,q,r,s}"
  shows "(p = r  q = s)  (q = r  p = s)"
proof -
  from proj2_set_Col {p,q,r,s}
  obtain l where "proj2_incident p l" and "proj2_incident q l"
    and "proj2_incident r l" and "proj2_incident s l"
    by (unfold proj2_set_Col_def) auto
  with r  s and p  S and q  S and r  S and s  S
  have "p = r  p = s" and "q = r  q = s"
    by (simp_all add: line_S_two_intersections_only)

  show "(p = r  q = s)  (q = r  p = s)"
  proof cases
    assume "p = r"
    with p  q and q = r  q = s
    show "(p = r  q = s)  (q = r  p = s)" by simp
  next
    assume "p  r"
    with p = r  p = s have "p = s" by simp
    with p  q and q = r  q = s
    show "(p = r  q = s)  (q = r  p = s)" by simp
  qed
qed

definition are_endpoints_in_S :: "[proj2, proj2, proj2, proj2]  bool" where
  "are_endpoints_in_S p q a b
   p  q  p  S  q  S  a  hyp2  b  hyp2  proj2_set_Col {p,q,a,b}"

lemma are_endpoints_in_S':
  assumes "p  q" and "a  b" and "p  S" and "q  S" and "a  hyp2  S"
  and "b  hyp2  S" and "proj2_set_Col {p,q,a,b}"
  shows "(p = endpoint_in_S a b  q = endpoint_in_S b a)
   (q = endpoint_in_S a b  p = endpoint_in_S b a)"
  (is "(p = ?r  q = ?s)  (q = ?r  p = ?s)")
proof -
  from a  b and a  hyp2  S and b  hyp2  S
  have "?r  ?s" by (simp add: endpoint_in_S_swap)

  from a  hyp2  S and b  hyp2  S
  have "?r  S" and "?s  S" by (simp_all add: endpoint_in_S)

  from proj2_set_Col {p,q,a,b}
  obtain l where "proj2_incident p l" and "proj2_incident q l"
    and "proj2_incident a l" and "proj2_incident b l"
    by (unfold proj2_set_Col_def) auto

  from a  b and a  hyp2  S and b  hyp2  S and proj2_incident a l
    and proj2_incident b l
  have "proj2_incident ?r l" and "proj2_incident ?s l"
    by (simp_all add: endpoint_in_S_incident)
  with proj2_incident p l and proj2_incident q l
  have "proj2_set_Col {p,q,?r,?s}"
    by (unfold proj2_set_Col_def) (simp add: exI [of _ l])
  with p  q and ?r  ?s and p  S and q  S and ?r  S and ?s  S
  show "(p = ?r  q = ?s)  (q = ?r  p = ?s)"
    by (rule line_S_match_intersections)
qed

lemma are_endpoints_in_S:
  assumes "a  b" and "are_endpoints_in_S p q a b"
  shows "(p = endpoint_in_S a b  q = endpoint_in_S b a)
   (q = endpoint_in_S a b  p = endpoint_in_S b a)"
  using assms
  by (unfold are_endpoints_in_S_def) (simp add: are_endpoints_in_S')

lemma S_intersections_endpoints_in_S:
  assumes "a  0" and "b  0" and "proj2_abs a  proj2_abs b" (is "?pa  ?pb")
  and "proj2_abs a  hyp2" and "proj2_abs b  hyp2  S"
  shows "(S_intersection1 a b = endpoint_in_S ?pa ?pb
       S_intersection2 a b = endpoint_in_S ?pb ?pa)
     (S_intersection2 a b = endpoint_in_S ?pa ?pb
       S_intersection1 a b = endpoint_in_S ?pb ?pa)"
  (is "(?pp = ?pr  ?pq = ?ps)  (?pq = ?pr  ?pp = ?ps)")
proof -
  from a  0 and b  0 and ?pa  ?pb and ?pa  hyp2
  have "?pp  ?pq" by (simp add: S_intersections_distinct)

  from a  0 and b  0 and ?pa  ?pb and proj2_abs a  hyp2
  have "?pp  S" and "?pq  S"
    by (simp_all add: S_intersections_in_S)

  let ?l = "proj2_line_through ?pa ?pb"
  have "proj2_incident ?pa ?l" and "proj2_incident ?pb ?l"
    by (rule proj2_line_through_incident)+
  with a  0 and b  0 and ?pa  ?pb
  have "proj2_incident ?pp ?l" and "proj2_incident ?pq ?l"
    by (rule S_intersections_incident)+
  with proj2_incident ?pa ?l and proj2_incident ?pb ?l
  have "proj2_set_Col {?pp,?pq,?pa,?pb}"
    by (unfold proj2_set_Col_def) (simp add: exI [of _ ?l])
  with ?pp  ?pq and ?pa  ?pb and ?pp  S and ?pq  S and ?pa  hyp2
    and ?pb  hyp2  S
  show "(?pp = ?pr  ?pq = ?ps)  (?pq = ?pr  ?pp = ?ps)"
    by (simp add: are_endpoints_in_S')
qed

lemma between_endpoints_in_S:
  assumes "a  b" and "a  hyp2  S" and "b  hyp2  S"
  shows "B
  (cart2_pt (endpoint_in_S a b)) (cart2_pt a) (cart2_pt (endpoint_in_S b a))"
  (is "B ?cp ?ca ?cq")
proof -
  let ?cb = "cart2_pt b"
  from b  hyp2  S and a  hyp2  S and a  b
  have "?cb  ?ca" by (auto simp add: hyp2_S_cart2_inj)

  from a  hyp2  S and b  hyp2  S
  have "B ?ca ?cb ?cp" and "B ?cb ?ca ?cq" by (simp_all add: endpoint_in_S)

  from B ?ca ?cb ?cp have "B ?cp ?cb ?ca" by (rule real_euclid.th3_2)
  with ?cb  ?ca and B ?cb ?ca ?cq
  show "B ?cp ?ca ?cq" by (simp add: real_euclid.th3_7_1)
qed

lemma S_hyp2_S_cart2_append1:
  assumes "p  q" and "p  S" and "q  S" and "a  hyp2"
  and "proj2_incident p l" and "proj2_incident q l" and "proj2_incident a l"
  shows " k. k > 0  k < 1
   cart2_append1 a = k *R cart2_append1 q + (1 - k) *R cart2_append1 p"
proof -
  from p  S and q  S and a  hyp2
  have "z_non_zero p" and "z_non_zero q" and "z_non_zero a"
    by (simp_all add: hyp2_S_z_non_zero)

  from assms
  have "B (cart2_pt p) (cart2_pt a) (cart2_pt q)" (is "B ?cp ?ca ?cq")
    by (simp add: hyp2_incident_in_middle)

  from p  S and q  S and a  hyp2
  have "a  p" and "a  q" by (simp_all add: hyp2_S_not_equal)

  with z_non_zero p and z_non_zero a and z_non_zero q
    and B ?cp ?ca ?cq
  show " k. k > 0  k < 1
     cart2_append1 a = k *R cart2_append1 q + (1 - k) *R cart2_append1 p"
    by (rule cart2_append1_between_strict)
qed

lemma are_endpoints_in_S_swap_34:
  assumes "are_endpoints_in_S p q a b"
  shows "are_endpoints_in_S p q b a"
proof -
  have "{p,q,b,a} = {p,q,a,b}" by auto
  with are_endpoints_in_S p q a b
  show "are_endpoints_in_S p q b a" by (unfold are_endpoints_in_S_def) simp
qed

lemma proj2_set_Col_endpoints_in_S:
  assumes "a  b" and "a  hyp2  S" and "b  hyp2  S"
  shows "proj2_set_Col {endpoint_in_S a b, endpoint_in_S b a, a, b}"
  (is "proj2_set_Col {?p,?q,a,b}")
proof -
  let ?l = "proj2_line_through a b"
  have "proj2_incident a ?l" and "proj2_incident b ?l"
    by (rule proj2_line_through_incident)+
  with a  b and a  hyp2  S and b  hyp2  S
  have "proj2_incident ?p ?l" and "proj2_incident ?q ?l"
    by (simp_all add: endpoint_in_S_incident)
  with proj2_incident a ?l and proj2_incident b ?l
  show "proj2_set_Col {?p,?q,a,b}"
    by (unfold proj2_set_Col_def) (simp add: exI [of _ ?l])
qed

lemma endpoints_in_S_are_endpoints_in_S:
  assumes "a  b" and "a  hyp2" and "b  hyp2"
  shows "are_endpoints_in_S (endpoint_in_S a b) (endpoint_in_S b a) a b"
  (is "are_endpoints_in_S ?p ?q a b")
proof -
  from a  b and a  hyp2 and b  hyp2
  have "?p  ?q" by (simp add: endpoint_in_S_swap)

  from a  hyp2 and b  hyp2
  have "?p  S" and "?q  S" by (simp_all add: endpoint_in_S)

  from assms
  have "proj2_set_Col {?p,?q,a,b}" by (simp add: proj2_set_Col_endpoints_in_S)
  with ?p  ?q and ?p  S and ?q  S and a  hyp2 and b  hyp2
  show "are_endpoints_in_S ?p ?q a b" by (unfold are_endpoints_in_S_def) simp
qed

lemma endpoint_in_S_S_hyp2_distinct:
  assumes "p  S" and "a  hyp2  S" and "p  a"
  shows "endpoint_in_S p a  p"
proof
  from p  a and p  S and a  hyp2  S
  have "B (cart2_pt p) (cart2_pt a) (cart2_pt (endpoint_in_S p a))"
    by (simp add: endpoint_in_S)

  assume "endpoint_in_S p a = p"
  with B (cart2_pt p) (cart2_pt a) (cart2_pt (endpoint_in_S p a))
  have "cart2_pt p = cart2_pt a" by (simp add: real_euclid.A6')
  with p  S and a  hyp2  S have "p = a" by (simp add: hyp2_S_cart2_inj)
  with p  a show False ..
qed

lemma endpoint_in_S_S_strict_hyp2_distinct:
  assumes "p  S" and "a  hyp2"
  shows "endpoint_in_S p a  p"
proof -
  from a  hyp2 and p  S
  have "p  a" by (rule hyp2_S_not_equal [symmetric])
  with assms
  show "endpoint_in_S p a  p" by (simp add: endpoint_in_S_S_hyp2_distinct)
qed

lemma end_and_opposite_are_endpoints_in_S:
  assumes "a  hyp2" and "b  hyp2" and "p  S"
  and "proj2_incident a l" and "proj2_incident b l" and "proj2_incident p l"
  shows "are_endpoints_in_S p (endpoint_in_S p b) a b"
  (is "are_endpoints_in_S p ?q a b")
proof -
  from p  S and b  hyp2
  have "p  ?q" by (rule endpoint_in_S_S_strict_hyp2_distinct [symmetric])

  from p  S and b  hyp2 have "?q  S" by (simp add: endpoint_in_S)

  from b  hyp2 and p  S
  have "p  b" by (rule hyp2_S_not_equal [symmetric])
  with p  S and b  hyp2 and proj2_incident p l and proj2_incident b l
  have "proj2_incident ?q l" by (simp add: endpoint_in_S_incident)
  with proj2_incident p l and proj2_incident a l and proj2_incident b l
  have "proj2_set_Col {p,?q,a,b}"
    by (unfold proj2_set_Col_def) (simp add: exI [of _ l])
  with p  ?q and p  S and ?q  S and a  hyp2 and b  hyp2
  show "are_endpoints_in_S p ?q a b" by (unfold are_endpoints_in_S_def) simp
qed

lemma real_hyp2_B_hyp2_cltn2:
  assumes "is_K2_isometry J" and "BK a b c"
  shows "BK (hyp2_cltn2 a J) (hyp2_cltn2 b J) (hyp2_cltn2 c J)"
  (is "BK ?aJ ?bJ ?cJ")
proof -
  from BK a b c
  have "B (hyp2_rep a) (hyp2_rep b) (hyp2_rep c)" by (unfold real_hyp2_B_def)
  with is_K2_isometry J
  have "B (cart2_pt (apply_cltn2 (Rep_hyp2 a) J))
    (cart2_pt (apply_cltn2 (Rep_hyp2 b) J))
    (cart2_pt (apply_cltn2 (Rep_hyp2 c) J))"
    by (unfold hyp2_rep_def) (simp add: Rep_hyp2 statement_63)
  moreover from is_K2_isometry J
  have "apply_cltn2 (Rep_hyp2 a) J  hyp2"
    and "apply_cltn2 (Rep_hyp2 b) J  hyp2"
    and "apply_cltn2 (Rep_hyp2 c) J  hyp2"
    by (rule apply_cltn2_Rep_hyp2)+
  ultimately show "BK (hyp2_cltn2 a J) (hyp2_cltn2 b J) (hyp2_cltn2 c J)"
    unfolding hyp2_cltn2_def and real_hyp2_B_def and hyp2_rep_def
    by (simp add: Abs_hyp2_inverse)
qed

lemma real_hyp2_C_hyp2_cltn2:
  assumes "is_K2_isometry J"
  shows "a b \<congruent>K (hyp2_cltn2 a J) (hyp2_cltn2 b J)" (is "a b \<congruent>K ?aJ ?bJ")
  using assms by (unfold real_hyp2_C_def) (simp add: exI [of _ J])

subsection ‹Perpendicularity›

definition M_perp :: "proj2_line  proj2_line  bool" where
  "M_perp l m  proj2_incident (pole l) m"

lemma M_perp_sym:
  assumes "M_perp l m"
  shows "M_perp m l"
proof -
  from M_perp l m have "proj2_incident (pole l) m" by (unfold M_perp_def)
  hence "proj2_incident (pole m) (polar (pole l))" by (rule incident_pole_polar)
  hence "proj2_incident (pole m) l" by (simp add: polar_pole)
  thus "M_perp m l" by (unfold M_perp_def)
qed

lemma M_perp_to_compass:
  assumes "M_perp l m" and "a  hyp2" and "proj2_incident a l"
  and "b  hyp2" and "proj2_incident b m"
  shows " J. is_K2_isometry J
   apply_cltn2_line equator J = l  apply_cltn2_line meridian J = m"
proof -
  from a  K2 and proj2_incident a l
    and line_through_K2_intersect_S_twice [of a l]
  obtain p and q where "p  q" and "p  S" and "q  S"
    and "proj2_incident p l" and "proj2_incident q l"
    by auto

  have " r. r  S  r  {p,q}  proj2_incident r m"
  proof cases
    assume "proj2_incident p m"

    from b  K2 and proj2_incident b m
      and line_through_K2_intersect_S_again [of b m]
    obtain r where "r  S" and "r  p" and "proj2_incident r m" by auto

    have "r  {p,q}"
    proof
      assume "r  {p,q}"
      with r  p have "r = q" by simp
      with proj2_incident r m have "proj2_incident q m" by simp
      with proj2_incident p l and proj2_incident q l
        and proj2_incident p m and proj2_incident q m and p  q
        and proj2_incident_unique [of p l q m]
      have "l = m" by simp
      with M_perp l m have "M_perp l l" by simp
      hence "proj2_incident (pole l) l" (is "proj2_incident ?s l")
        by (unfold M_perp_def)
      hence "proj2_incident ?s (polar ?s)" by (subst polar_pole)
      hence "?s  S" by (simp add: incident_own_polar_in_S)
      with p  S and q  S and proj2_incident p l and proj2_incident q l
        and point_in_S_polar_is_tangent [of ?s]
      have "p = ?s" and "q = ?s" by (auto simp add: polar_pole)
      with p  q show False by simp
    qed
    with r  S and proj2_incident r m
    show " r. r  S  r  {p,q}  proj2_incident r m"
      by (simp add: exI [of _ r])
  next
    assume "¬ proj2_incident p m"

    from b  K2 and proj2_incident b m
      and line_through_K2_intersect_S_again [of b m]
    obtain r where "r  S" and "r  q" and "proj2_incident r m" by auto

    from ¬ proj2_incident p m and proj2_incident r m have "r  p" by auto
    with r  S and r  q and proj2_incident r m
    show " r. r  S  r  {p,q}  proj2_incident r m"
      by (simp add: exI [of _ r])
  qed
  then obtain r where "r  S" and "r  {p,q}" and "proj2_incident r m" by auto

  from p  S and q  S and r  S and p  q and r  {p,q}
    and statement65_special_case [of p q r]
  obtain J where "is_K2_isometry J" and "apply_cltn2 east J = p"
    and "apply_cltn2 west J = q" and "apply_cltn2 north J = r"
    and "apply_cltn2 far_north J = proj2_intersection (polar p) (polar q)"
    by auto

  from apply_cltn2 east J = p and apply_cltn2 west J = q
    and proj2_incident p l and proj2_incident q l
  have "proj2_incident (apply_cltn2 east J) l"
    and "proj2_incident (apply_cltn2 west J) l"
    by simp_all
  with east_west_distinct and east_west_on_equator
  have "apply_cltn2_line equator J = l" by (rule apply_cltn2_line_unique)

  from apply_cltn2 north J = r and proj2_incident r m
  have "proj2_incident (apply_cltn2 north J) m" by simp

  from p  q and polar_inj have "polar p  polar q" by fast

  from proj2_incident p l and proj2_incident q l
  have "proj2_incident (pole l) (polar p)"
    and "proj2_incident (pole l) (polar q)"
    by (simp_all add: incident_pole_polar)
  with polar p  polar q
  have "pole l = proj2_intersection (polar p) (polar q)"
    by (rule proj2_intersection_unique)
  with apply_cltn2 far_north J = proj2_intersection (polar p) (polar q)
  have "apply_cltn2 far_north J = pole l" by simp
  with M_perp l m
  have "proj2_incident (apply_cltn2 far_north J) m" by (unfold M_perp_def) simp
  with north_far_north_distinct and north_south_far_north_on_meridian
    and proj2_incident (apply_cltn2 north J) m
  have "apply_cltn2_line meridian J = m" by (simp add: apply_cltn2_line_unique)
  with is_K2_isometry J and apply_cltn2_line equator J = l
  show " J. is_K2_isometry J
     apply_cltn2_line equator J = l  apply_cltn2_line meridian J = m"
    by (simp add: exI [of _ J])
qed

definition drop_perp :: "proj2  proj2_line  proj2_line" where
  "drop_perp p l  proj2_line_through p (pole l)"

lemma drop_perp_incident: "proj2_incident p (drop_perp p l)"
  by (unfold drop_perp_def) (rule proj2_line_through_incident)

lemma drop_perp_perp: "M_perp l (drop_perp p l)"
  by (unfold drop_perp_def M_perp_def) (rule proj2_line_through_incident)

definition perp_foot :: "proj2  proj2_line  proj2" where
  "perp_foot p l  proj2_intersection l (drop_perp p l)"

lemma perp_foot_incident:
  shows "proj2_incident (perp_foot p l) l"
  and "proj2_incident (perp_foot p l) (drop_perp p l)"
  by (unfold perp_foot_def) (rule proj2_intersection_incident)+

lemma M_perp_hyp2:
  assumes "M_perp l m" and "a  hyp2" and "proj2_incident a l" and "b  hyp2"
  and "proj2_incident b m" and "proj2_incident c l" and "proj2_incident c m"
  shows "c  hyp2"
proof -
  from M_perp l m and a  hyp2 and proj2_incident a l and b  hyp2
    and proj2_incident b m and M_perp_to_compass [of l m a b]
  obtain J where "is_K2_isometry J" and "apply_cltn2_line equator J = l"
    and "apply_cltn2_line meridian J = m"
    by auto

  from is_K2_isometry J and K2_centre_in_K2
  have "apply_cltn2 K2_centre J  hyp2"
    by (rule statement60_one_way)

  from proj2_incident c l and apply_cltn2_line equator J = l
    and proj2_incident c m and apply_cltn2_line meridian J = m
  have "proj2_incident c (apply_cltn2_line equator J)"
    and "proj2_incident c (apply_cltn2_line meridian J)"
    by simp_all
  with equator_meridian_distinct and K2_centre_on_equator_meridian
  have "apply_cltn2 K2_centre J = c" by (rule apply_cltn2_unique)
  with apply_cltn2 K2_centre J  hyp2 show "c  hyp2" by simp
qed

lemma perp_foot_hyp2:
  assumes "a  hyp2" and "proj2_incident a l" and "b  hyp2"
  shows "perp_foot b l  hyp2"
  using drop_perp_perp [of l b] and a  hyp2 and proj2_incident a l
    and b  hyp2 and drop_perp_incident [of b l]
    and perp_foot_incident [of b l]
  by (rule M_perp_hyp2)

definition perp_up :: "proj2  proj2_line  proj2" where
  "perp_up a l
   if proj2_incident a l then ϵ p. p  S  proj2_incident p (drop_perp a l)
  else endpoint_in_S (perp_foot a l) a"

lemma perp_up_degenerate_in_S_incident:
  assumes "a  hyp2" and "proj2_incident a l"
  shows "perp_up a l  S" (is "?p  S")
  and "proj2_incident (perp_up a l) (drop_perp a l)"
proof -
  from proj2_incident a l
  have "?p = (ϵ p. p  S  proj2_incident p (drop_perp a l))"
    by (unfold perp_up_def) simp

  from a  hyp2 and drop_perp_incident [of a l]
  have " p. p  S  proj2_incident p (drop_perp a l)"
    by (rule line_through_K2_intersect_S)
  hence "?p  S  proj2_incident ?p (drop_perp a l)"
    unfolding ?p = (ϵ p. p  S  proj2_incident p (drop_perp a l))
    by (rule someI_ex)
  thus "?p  S" and "proj2_incident ?p (drop_perp a l)" by simp_all
qed

lemma perp_up_non_degenerate_in_S_at_end:
  assumes "a  hyp2" and "b  hyp2" and "proj2_incident b l"
  and "¬ proj2_incident a l"
  shows "perp_up a l  S"
  and "B (cart2_pt (perp_foot a l)) (cart2_pt a) (cart2_pt (perp_up a l))"
proof -
  from ¬ proj2_incident a l
  have "perp_up a l = endpoint_in_S (perp_foot a l) a"
    by (unfold perp_up_def) simp

  from b  hyp2 and proj2_incident b l and a  hyp2
  have "perp_foot a l  hyp2" by (rule perp_foot_hyp2)
  with a  hyp2
  show "perp_up a l  S"
    and "B (cart2_pt (perp_foot a l)) (cart2_pt a) (cart2_pt (perp_up a l))"
    unfolding perp_up a l = endpoint_in_S (perp_foot a l) a
    by (simp_all add: endpoint_in_S)
qed

lemma perp_up_in_S:
  assumes "a  hyp2" and "b  hyp2" and "proj2_incident b l"
  shows "perp_up a l  S"
proof cases
  assume "proj2_incident a l"
  with a  hyp2
  show "perp_up a l  S" by (rule perp_up_degenerate_in_S_incident)
next
  assume "¬ proj2_incident a l"
  with assms
  show "perp_up a l  S" by (rule perp_up_non_degenerate_in_S_at_end)
qed

lemma perp_up_incident:
  assumes "a  hyp2" and "b  hyp2" and "proj2_incident b l"
  shows "proj2_incident (perp_up a l) (drop_perp a l)"
  (is "proj2_incident ?p ?m")
proof cases
  assume "proj2_incident a l"
  with a  hyp2
  show "proj2_incident ?p ?m" by (rule perp_up_degenerate_in_S_incident)
next
  assume "¬ proj2_incident a l"
  hence "?p = endpoint_in_S (perp_foot a l) a" (is "?p = endpoint_in_S ?c a")
    by (unfold perp_up_def) simp

  from perp_foot_incident [of a l] and ¬ proj2_incident a l
  have "?c  a" by auto

  from b  hyp2 and proj2_incident b l and a  hyp2
  have "?c  hyp2" by (rule perp_foot_hyp2)
  with ?c  a and a  hyp2 and drop_perp_incident [of a l]
    and perp_foot_incident [of a l]
  show "proj2_incident ?p ?m"
    by (unfold ?p = endpoint_in_S ?c a) (simp add: endpoint_in_S_incident)
qed

lemma drop_perp_same_line_pole_in_S:
  assumes "drop_perp p l = l"
  shows "pole l  S"
proof -
  from drop_perp p l = l
  have "l = proj2_line_through p (pole l)" by (unfold drop_perp_def) simp
  with proj2_line_through_incident [of "pole l" p]
  have "proj2_incident (pole l) l" by simp
  hence "proj2_incident (pole l) (polar (pole l))" by (subst polar_pole)
  thus "pole l  S" by (unfold incident_own_polar_in_S)
qed

lemma hyp2_drop_perp_not_same_line:
  assumes "a  hyp2"
  shows "drop_perp a l  l"
proof
  assume "drop_perp a l = l"
  hence "pole l  S" by (rule drop_perp_same_line_pole_in_S)
  with a  hyp2
  have "¬ proj2_incident a (polar (pole l))"
    by (simp add: tangent_not_through_K2)
  with drop_perp a l = l
  have "¬ proj2_incident a (drop_perp a l)" by (simp add: polar_pole)
  with drop_perp_incident [of a l] show False by simp
qed

lemma hyp2_incident_perp_foot_same_point:
  assumes "a  hyp2" and "proj2_incident a l"
  shows "perp_foot a l = a"
proof -
  from a  hyp2
  have "drop_perp a l  l" by (rule hyp2_drop_perp_not_same_line)
  with perp_foot_incident [of a l] and proj2_incident a l
    and drop_perp_incident [of a l] and proj2_incident_unique
  show "perp_foot a l = a" by fast
qed

lemma perp_up_at_end:
  assumes "a  hyp2" and "b  hyp2" and "proj2_incident b l"
  shows "B (cart2_pt (perp_foot a l)) (cart2_pt a) (cart2_pt (perp_up a l))"
proof cases
  assume "proj2_incident a l"
  with a  hyp2
  have "perp_foot a l = a" by (rule hyp2_incident_perp_foot_same_point)
  thus "B (cart2_pt (perp_foot a l)) (cart2_pt a) (cart2_pt (perp_up a l))"
    by (simp add: real_euclid.th3_1 real_euclid.th3_2)
next
  assume "¬ proj2_incident a l"
  with assms
  show "B (cart2_pt (perp_foot a l)) (cart2_pt a) (cart2_pt (perp_up a l))"
    by (rule perp_up_non_degenerate_in_S_at_end)
qed

definition perp_down :: "proj2  proj2_line  proj2" where
  "perp_down a l  endpoint_in_S (perp_up a l) a"

lemma perp_down_in_S:
  assumes "a  hyp2" and "b  hyp2" and "proj2_incident b l"
  shows "perp_down a l  S"
proof -
  from assms have "perp_up a l  S" by (rule perp_up_in_S)
  with a  hyp2
  show "perp_down a l  S" by (unfold perp_down_def) (simp add: endpoint_in_S)
qed

lemma perp_down_incident:
  assumes "a  hyp2" and "b  hyp2" and "proj2_incident b l"
  shows "proj2_incident (perp_down a l) (drop_perp a l)"
proof -
  from assms have "perp_up a l  S" by (rule perp_up_in_S)
  with a  hyp2 have "perp_up a l  a" by (rule hyp2_S_not_equal [symmetric])

  from assms
  have "proj2_incident (perp_up a l) (drop_perp a l)" by (rule perp_up_incident)
  with perp_up a l  a and perp_up a l  S and a  hyp2
    and drop_perp_incident [of a l]
  show "proj2_incident (perp_down a l) (drop_perp a l)"
    by (unfold perp_down_def) (simp add: endpoint_in_S_incident)
qed

lemma perp_up_down_distinct:
  assumes "a  hyp2" and "b  hyp2" and "proj2_incident b l"
  shows "perp_up a l  perp_down a l"
proof -
  from assms have "perp_up a l  S" by (rule perp_up_in_S)
  with a  hyp2
  show "perp_up a l  perp_down a l"
    unfolding perp_down_def
    by (simp add: endpoint_in_S_S_strict_hyp2_distinct [symmetric])
qed

lemma perp_up_down_foot_are_endpoints_in_S:
  assumes "a  hyp2" and "b  hyp2" and "proj2_incident b l"
  shows "are_endpoints_in_S (perp_up a l) (perp_down a l) (perp_foot a l) a"
proof -
  from b  hyp2 and proj2_incident b l and a  hyp2
  have "perp_foot a l  hyp2" by (rule perp_foot_hyp2)

  from assms have "perp_up a l  S" by (rule perp_up_in_S)

  from assms
  have "proj2_incident (perp_up a l) (drop_perp a l)" by (rule perp_up_incident)
  with perp_foot a l  hyp2 and a  hyp2 and perp_up a l  S
    and perp_foot_incident(2) [of a l] and drop_perp_incident [of a l]
  show "are_endpoints_in_S (perp_up a l) (perp_down a l) (perp_foot a l) a"
    by (unfold perp_down_def) (rule end_and_opposite_are_endpoints_in_S)
qed

lemma perp_foot_opposite_endpoint_in_S:
  assumes "a  hyp2" and "b  hyp2" and "c  hyp2" and "a  b"
  shows
  "endpoint_in_S (endpoint_in_S a b) (perp_foot c (proj2_line_through a b))
  = endpoint_in_S b a"
  (is "endpoint_in_S ?p ?d = endpoint_in_S b a")
proof -
  let ?q = "endpoint_in_S ?p ?d"

  from a  hyp2 and b  hyp2 have "?p  S" by (simp add: endpoint_in_S)

  let ?l = "proj2_line_through a b"
  have "proj2_incident a ?l" and "proj2_incident b ?l"
    by (rule proj2_line_through_incident)+
  with a  b and a  hyp2 and b  hyp2
  have "proj2_incident ?p ?l"
    by (simp_all add: endpoint_in_S_incident)

  from a  hyp2 and proj2_incident a ?l and c  hyp2
  have "?d  hyp2" by (rule perp_foot_hyp2)
  with ?p  S have "?q  ?p" by (rule endpoint_in_S_S_strict_hyp2_distinct)

  from ?p  S and ?d  hyp2 have "?q  S" by (simp add: endpoint_in_S)

  from ?d  hyp2 and ?p  S
  have "?p  ?d" by (rule hyp2_S_not_equal [symmetric])
  with ?p  S and ?d  hyp2 and proj2_incident ?p ?l
    and perp_foot_incident(1) [of c ?l]
  have "proj2_incident ?q ?l" by (simp add: endpoint_in_S_incident)
  with a  b and a  hyp2 and b  hyp2 and ?q  S
    and proj2_incident a ?l and proj2_incident b ?l
  have "?q = ?p  ?q = endpoint_in_S b a"
    by (simp add: endpoints_in_S_incident_unique)
  with ?q  ?p show "?q = endpoint_in_S b a" by simp
qed

lemma endpoints_in_S_perp_foot_are_endpoints_in_S:
  assumes "a  hyp2" and "b  hyp2" and "c  hyp2" and "a  b"
  and "proj2_incident a l" and "proj2_incident b l"
  shows "are_endpoints_in_S
  (endpoint_in_S a b) (endpoint_in_S b a) a (perp_foot c l)"
proof -
  define p q d
    where "p = endpoint_in_S a b"
      and "q = endpoint_in_S b a"
      and "d = perp_foot c l"

  from a  b and a  hyp2 and b  hyp2
  have "p  q" by (unfold p_def q_def) (simp add: endpoint_in_S_swap)

  from a  hyp2 and b  hyp2
  have "p  S" and "q  S" by (unfold p_def q_def) (simp_all add: endpoint_in_S)

  from a  hyp2 and proj2_incident a l and c  hyp2
  have "d  hyp2" by (unfold d_def) (rule perp_foot_hyp2)

  from a  b and a  hyp2 and b  hyp2 and proj2_incident a l
    and proj2_incident b l
  have "proj2_incident p l" and "proj2_incident q l"
    by (unfold p_def q_def) (simp_all add: endpoint_in_S_incident)
  with proj2_incident a l and perp_foot_incident(1) [of c l]
  have "proj2_set_Col {p,q,a,d}"
    by (unfold d_def proj2_set_Col_def) (simp add: exI [of _ l])
  with p  q and p  S and q  S and a  hyp2 and d  hyp2
  show "are_endpoints_in_S p q a d" by (unfold are_endpoints_in_S_def) simp
qed

definition right_angle :: "proj2  proj2  proj2  bool" where
  "right_angle p a q
   p  S  q  S  a  hyp2
   M_perp (proj2_line_through p a) (proj2_line_through a q)"

lemma perp_foot_up_right_angle:
  assumes "p  S" and "a  hyp2" and "b  hyp2" and "proj2_incident p l"
  and "proj2_incident b l"
  shows "right_angle p (perp_foot a l) (perp_up a l)"
proof -
  define c where "c = perp_foot a l"
  define q where "q = perp_up a l"
  from a  hyp2 and b  hyp2 and proj2_incident b l
  have "q  S" by (unfold q_def) (rule perp_up_in_S)

  from b  hyp2 and proj2_incident b l and a  hyp2
  have "c  hyp2" by (unfold c_def) (rule perp_foot_hyp2)
  with p  S and q  S have "c  p" and "c  q"
    by (simp_all add: hyp2_S_not_equal)

  from c  p [symmetric] and proj2_incident p l
    and perp_foot_incident(1) [of a l]
  have "l = proj2_line_through p c"
    by (unfold c_def) (rule proj2_line_through_unique)

  define m where "m = drop_perp a l"
  from a  hyp2 and b  hyp2 and proj2_incident b l
  have "proj2_incident q m" by (unfold q_def m_def) (rule perp_up_incident)
  with c  q and perp_foot_incident(2) [of a l]
  have "m = proj2_line_through c q"
    by (unfold c_def m_def) (rule proj2_line_through_unique)
  with p  S and q  S and c  hyp2 and drop_perp_perp [of l a]
    and l = proj2_line_through p c
  show "right_angle p (perp_foot a l) (perp_up a l)"
    by (unfold right_angle_def q_def c_def m_def) simp
qed

lemma M_perp_unique:
  assumes "a  hyp2" and "b  hyp2" and "proj2_incident a l"
  and "proj2_incident b m" and "proj2_incident b n" and "M_perp l m"
  and "M_perp l n"
  shows "m = n"
proof -
  from a  hyp2 and proj2_incident a l
  have "pole l  hyp2" by (rule line_through_hyp2_pole_not_in_hyp2)
  with b  hyp2 have "b  pole l" by auto
  with proj2_incident b m and M_perp l m and proj2_incident b n
    and M_perp l n and proj2_incident_unique
  show "m = n" by (unfold M_perp_def) auto
qed

lemma perp_foot_eq_implies_drop_perp_eq:
  assumes "a  hyp2" and "b  hyp2" and "proj2_incident a l"
  and "perp_foot b l = perp_foot c l"
  shows "drop_perp b l = drop_perp c l"
proof -
  from a  hyp2 and proj2_incident a l and b  hyp2
  have "perp_foot b l  hyp2" by (rule perp_foot_hyp2)

  from perp_foot b l = perp_foot c l
  have "proj2_incident (perp_foot b l) (drop_perp c l)"
    by (simp add: perp_foot_incident)
  with a  hyp2 and perp_foot b l  hyp2 and proj2_incident a l
    and perp_foot_incident(2) [of b l] and drop_perp_perp [of l]
  show "drop_perp b l = drop_perp c l" by (simp add: M_perp_unique)
qed

lemma right_angle_to_compass:
  assumes "right_angle p a q"
  shows " J. is_K2_isometry J  apply_cltn2 p J = east
   apply_cltn2 a J = K2_centre  apply_cltn2 q J = north"
proof -
  from right_angle p a q
  have "p  S" and "q  S" and "a  hyp2"
    and "M_perp (proj2_line_through p a) (proj2_line_through a q)"
    (is "M_perp ?l ?m")
    by (unfold right_angle_def) simp_all

  have "proj2_incident p ?l" and "proj2_incident a ?l"
    and "proj2_incident q ?m" and "proj2_incident a ?m"
    by (rule proj2_line_through_incident)+

  from M_perp ?l ?m and a  hyp2 and proj2_incident a ?l
    and proj2_incident a ?m and M_perp_to_compass [of ?l ?m a a]
  obtain J''i where "is_K2_isometry J''i"
    and "apply_cltn2_line equator J''i = ?l"
    and "apply_cltn2_line meridian J''i = ?m"
    by auto
  let ?J'' = "cltn2_inverse J''i"

  from apply_cltn2_line equator J''i = ?l
    and apply_cltn2_line meridian J''i = ?m
    and proj2_incident p ?l and proj2_incident a ?l
    and proj2_incident q ?m and proj2_incident a ?m
  have "proj2_incident (apply_cltn2 p ?J'') equator"
    and "proj2_incident (apply_cltn2 a ?J'') equator"
    and "proj2_incident (apply_cltn2 q ?J'') meridian"
    and "proj2_incident (apply_cltn2 a ?J'') meridian"
    by (simp_all add: apply_cltn2_incident [symmetric])

  from proj2_incident (apply_cltn2 a ?J'') equator
    and proj2_incident (apply_cltn2 a ?J'') meridian
  have "apply_cltn2 a ?J'' = K2_centre"
    by (rule on_equator_meridian_is_K2_centre)

  from is_K2_isometry J''i
  have "is_K2_isometry ?J''" by (rule cltn2_inverse_is_K2_isometry)
  with p  S and q  S
  have "apply_cltn2 p ?J''  S" and "apply_cltn2 q ?J''  S"
    by (unfold is_K2_isometry_def) simp_all
  with east_west_distinct and north_south_distinct and compass_in_S
    and east_west_on_equator and north_south_far_north_on_meridian
    and proj2_incident (apply_cltn2 p ?J'') equator
    and proj2_incident (apply_cltn2 q ?J'') meridian
  have "apply_cltn2 p ?J'' = east  apply_cltn2 p ?J'' = west"
    and "apply_cltn2 q ?J'' = north  apply_cltn2 q ?J'' = south"
    by (simp_all add: line_S_two_intersections_only)

  have " J'. is_K2_isometry J'  apply_cltn2 p J' = east
     apply_cltn2 a J' = K2_centre
     (apply_cltn2 q J' = north  apply_cltn2 q J' = south)"
  proof cases
    assume "apply_cltn2 p ?J'' = east"
    with is_K2_isometry ?J'' and apply_cltn2 a ?J'' = K2_centre
      and apply_cltn2 q ?J'' = north  apply_cltn2 q ?J'' = south
    show " J'. is_K2_isometry J'  apply_cltn2 p J' = east
       apply_cltn2 a J' = K2_centre
       (apply_cltn2 q J' = north  apply_cltn2 q J' = south)"
      by (simp add: exI [of _ ?J''])
  next
    assume "apply_cltn2 p ?J''  east"
    with apply_cltn2 p ?J'' = east  apply_cltn2 p ?J'' = west
    have "apply_cltn2 p ?J'' = west" by simp

    let ?J' = "cltn2_compose ?J'' meridian_reflect"
    from is_K2_isometry ?J'' and meridian_reflect_K2_isometry
    have "is_K2_isometry ?J'" by (rule cltn2_compose_is_K2_isometry)
    moreover
    from apply_cltn2 p ?J'' = west and apply_cltn2 a ?J'' = K2_centre
      and apply_cltn2 q ?J'' = north  apply_cltn2 q ?J'' = south
      and compass_reflect_compass
    have "apply_cltn2 p ?J' = east" and "apply_cltn2 a ?J' = K2_centre"
      and "apply_cltn2 q ?J' = north  apply_cltn2 q ?J' = south"
      by (auto simp add: cltn2.act_act [simplified, symmetric])
    ultimately
    show " J'. is_K2_isometry J'  apply_cltn2 p J' = east
       apply_cltn2 a J' = K2_centre
       (apply_cltn2 q J' = north  apply_cltn2 q J' = south)"
      by (simp add: exI [of _ ?J'])
  qed
  then obtain J' where "is_K2_isometry J'" and "apply_cltn2 p J' = east"
    and "apply_cltn2 a J' = K2_centre"
    and "apply_cltn2 q J' = north  apply_cltn2 q J' = south"
    by auto

  show " J. is_K2_isometry J  apply_cltn2 p J = east
     apply_cltn2 a J = K2_centre  apply_cltn2 q J = north"
  proof cases
    assume "apply_cltn2 q J' = north"
    with is_K2_isometry J' and apply_cltn2 p J' = east
      and apply_cltn2 a J' = K2_centre
    show " J. is_K2_isometry J  apply_cltn2 p J = east
       apply_cltn2 a J = K2_centre  apply_cltn2 q J = north"
      by (simp add: exI [of _ J'])
  next
    assume "apply_cltn2 q J'  north"
    with apply_cltn2 q J' = north  apply_cltn2 q J' = south
    have "apply_cltn2 q J' = south" by simp

    let ?J = "cltn2_compose J' equator_reflect"
    from is_K2_isometry J' and equator_reflect_K2_isometry
    have "is_K2_isometry ?J" by (rule cltn2_compose_is_K2_isometry)
    moreover
    from apply_cltn2 p J' = east and apply_cltn2 a J' = K2_centre
      and apply_cltn2 q J' = south and compass_reflect_compass
    have "apply_cltn2 p ?J = east" and "apply_cltn2 a ?J = K2_centre"
      and "apply_cltn2 q ?J = north"
      by (auto simp add: cltn2.act_act [simplified, symmetric])
    ultimately
    show " J. is_K2_isometry J  apply_cltn2 p J = east
       apply_cltn2 a J = K2_centre  apply_cltn2 q J = north"
      by (simp add: exI [of _ ?J])
  qed
qed

lemma right_angle_to_right_angle:
  assumes "right_angle p a q" and "right_angle r b s"
  shows " J. is_K2_isometry J
   apply_cltn2 p J = r  apply_cltn2 a J = b  apply_cltn2 q J = s"
proof -
  from right_angle p a q and right_angle_to_compass [of p a q]
  obtain H where "is_K2_isometry H" and "apply_cltn2 p H = east"
    and "apply_cltn2 a H = K2_centre" and "apply_cltn2 q H = north"
    by auto

  from right_angle r b s and right_angle_to_compass [of r b s]
  obtain K where "is_K2_isometry K" and "apply_cltn2 r K = east"
    and "apply_cltn2 b K = K2_centre" and "apply_cltn2 s K = north"
    by auto

  let ?Ki = "cltn2_inverse K"
  let ?J = "cltn2_compose H ?Ki"
  from is_K2_isometry H and is_K2_isometry K
  have "is_K2_isometry ?J"
    by (simp add: cltn2_inverse_is_K2_isometry cltn2_compose_is_K2_isometry)

  from apply_cltn2 r K = east and apply_cltn2 b K = K2_centre
    and apply_cltn2 s K = north
  have "apply_cltn2 east ?Ki = r" and "apply_cltn2 K2_centre ?Ki = b"
    and "apply_cltn2 north ?Ki = s"
    by (simp_all add: cltn2.act_inv_iff [simplified])
  with apply_cltn2 p H = east and apply_cltn2 a H = K2_centre
    and apply_cltn2 q H = north
  have "apply_cltn2 p ?J = r" and "apply_cltn2 a ?J = b"
    and "apply_cltn2 q ?J = s"
    by (simp_all add: cltn2.act_act [simplified,symmetric])
  with is_K2_isometry ?J
  show " J. is_K2_isometry J
     apply_cltn2 p J = r  apply_cltn2 a J = b  apply_cltn2 q J = s"
    by (simp add: exI [of _ ?J])
qed

subsection ‹Functions of distance›

definition exp_2dist :: "proj2  proj2  real" where
  "exp_2dist a b
   if a = b
  then 1
  else cross_ratio (endpoint_in_S a b) (endpoint_in_S b a) a b"

definition cosh_dist :: "proj2  proj2  real" where
  "cosh_dist a b  (sqrt (exp_2dist a b) + sqrt (1 / (exp_2dist a b))) / 2"

lemma exp_2dist_formula:
  assumes "a  0" and "b  0" and "proj2_abs a  hyp2" (is "?pa  hyp2")
  and "proj2_abs b  hyp2" (is "?pb  hyp2")
  shows "exp_2dist (proj2_abs a) (proj2_abs b)
    = (a  (M *v b) + sqrt (quarter_discrim a b))
      / (a  (M *v b) - sqrt (quarter_discrim a b))
   exp_2dist (proj2_abs a) (proj2_abs b)
    = (a  (M *v b) - sqrt (quarter_discrim a b))
      / (a  (M *v b) + sqrt (quarter_discrim a b))"
  (is "?e2d = (?aMb + ?sqd) / (?aMb - ?sqd)
     ?e2d = (?aMb - ?sqd) / (?aMb + ?sqd)")
proof cases
  assume "?pa = ?pb"
  hence "?e2d = 1" by (unfold exp_2dist_def, simp)

  from ?pa = ?pb
  have "quarter_discrim a b = 0" by (rule quarter_discrim_self_zero)
  hence "?sqd = 0" by simp

  from proj2_abs a = proj2_abs b and b  0 and proj2_abs_abs_mult
  obtain k where "a = k *R b" by auto

  from b  0 and proj2_abs b  hyp2
  have "b  (M *v b) < 0" by (subst K2_abs [symmetric])
  with a  0 and a = k *R b have "?aMb  0" by simp
  with ?e2d = 1 and ?sqd = 0
  show "?e2d = (?aMb + ?sqd) / (?aMb - ?sqd)
     ?e2d = (?aMb - ?sqd) / (?aMb + ?sqd)"
    by simp
next
  assume "?pa  ?pb"
  let ?l = "proj2_line_through ?pa ?pb"
  have "proj2_incident ?pa ?l" and "proj2_incident ?pb ?l"
    by (rule proj2_line_through_incident)+
  with a  0 and b  0 and ?pa  ?pb
  have "proj2_incident (S_intersection1 a b) ?l" (is "proj2_incident ?Si1 ?l")
    and "proj2_incident (S_intersection2 a b) ?l" (is "proj2_incident ?Si2 ?l")
    by (rule S_intersections_incident)+
  with proj2_incident ?pa ?l and proj2_incident ?pb ?l
  have "proj2_set_Col {?pa,?pb,?Si1,?Si2}" by (unfold proj2_set_Col_def, auto)

  have "{?pa,?pb,?Si2,?Si1} = {?pa,?pb,?Si1,?Si2}" by auto

  from a  0 and b  0 and ?pa  ?pb and ?pa  hyp2
  have "?Si1  S" and "?Si2  S"
    by (simp_all add: S_intersections_in_S)
  with ?pa  hyp2 and ?pb  hyp2
  have "?Si1  ?pa" and "?Si2  ?pa" and "?Si1  ?pb" and "?Si2  ?pb"
    by (simp_all add: hyp2_S_not_equal [symmetric])
  with proj2_set_Col {?pa,?pb,?Si1,?Si2} and ?pa  ?pb
  have "cross_ratio_correct ?pa ?pb ?Si1 ?Si2"
    and "cross_ratio_correct ?pa ?pb ?Si2 ?Si1"
    unfolding cross_ratio_correct_def
    by (simp_all add: {?pa,?pb,?Si2,?Si1} = {?pa,?pb,?Si1,?Si2})

  from a  0 and b  0 and ?pa  ?pb and ?pa  hyp2
  have "?Si1  ?Si2" by (simp add: S_intersections_distinct)
  with cross_ratio_correct ?pa ?pb ?Si1 ?Si2
    and cross_ratio_correct ?pa ?pb ?Si2 ?Si1
  have "cross_ratio ?Si1 ?Si2 ?pa ?pb = cross_ratio ?pa ?pb ?Si1 ?Si2"
    and "cross_ratio ?Si2 ?Si1 ?pa ?pb = cross_ratio ?pa ?pb ?Si2 ?Si1"
    by (simp_all add: cross_ratio_swap_13_24)

  from a  0 and proj2_abs a  hyp2
  have "a  (M *v a) < 0" by (subst K2_abs [symmetric])
  with a  0 and b  0 and ?pa  ?pb and cross_ratio_abs [of a b 1 1]
  have "cross_ratio ?pa ?pb ?Si1 ?Si2 = (-?aMb - ?sqd) / (-?aMb + ?sqd)"
    by (unfold S_intersections_defs S_intersection_coeffs_defs, simp)
  with times_divide_times_eq [of "-1" "-1" "-?aMb - ?sqd" "-?aMb + ?sqd"]
  have "cross_ratio ?pa ?pb ?Si1 ?Si2 = (?aMb + ?sqd) / (?aMb - ?sqd)" by (simp add: ac_simps)
  with cross_ratio ?Si1 ?Si2 ?pa ?pb = cross_ratio ?pa ?pb ?Si1 ?Si2
  have "cross_ratio ?Si1 ?Si2 ?pa ?pb = (?aMb + ?sqd) / (?aMb - ?sqd)" by simp

  from cross_ratio ?pa ?pb ?Si1 ?Si2 = (?aMb + ?sqd) / (?aMb - ?sqd)
    and cross_ratio_swap_34 [of ?pa ?pb ?Si2 ?Si1]
  have "cross_ratio ?pa ?pb ?Si2 ?Si1 = (?aMb - ?sqd) / (?aMb + ?sqd)" by simp
  with cross_ratio ?Si2 ?Si1 ?pa ?pb = cross_ratio ?pa ?pb ?Si2 ?Si1
  have "cross_ratio ?Si2 ?Si1 ?pa ?pb = (?aMb - ?sqd) / (?aMb + ?sqd)" by simp

  from a  0 and b  0 and ?pa  ?pb and ?pa  hyp2 and ?pb  hyp2
  have "(?Si1 = endpoint_in_S ?pa ?pb  ?Si2 = endpoint_in_S ?pb ?pa)
     (?Si2 = endpoint_in_S ?pa ?pb  ?Si1 = endpoint_in_S ?pb ?pa)"
    by (simp add: S_intersections_endpoints_in_S)
  with cross_ratio ?Si1 ?Si2 ?pa ?pb = (?aMb + ?sqd) / (?aMb - ?sqd)
    and cross_ratio ?Si2 ?Si1 ?pa ?pb = (?aMb - ?sqd) / (?aMb + ?sqd)
    and ?pa  ?pb
  show "?e2d = (?aMb + ?sqd) / (?aMb - ?sqd)
     ?e2d = (?aMb - ?sqd) / (?aMb + ?sqd)"
    by (unfold exp_2dist_def, auto)
qed

lemma cosh_dist_formula:
  assumes "a  0" and "b  0" and "proj2_abs a  hyp2" (is "?pa  hyp2")
  and "proj2_abs b  hyp2" (is "?pb  hyp2")
  shows "cosh_dist (proj2_abs a) (proj2_abs b)
  = ¦a  (M *v b)¦ / sqrt (a  (M *v a) * (b  (M *v b)))"
  (is "cosh_dist ?pa ?pb = ¦?aMb¦ / sqrt (?aMa * ?bMb)")
proof -
  let ?qd = "quarter_discrim a b"
  let ?sqd = "sqrt ?qd"
  let ?e2d = "exp_2dist ?pa ?pb"
  from assms
  have "?e2d = (?aMb + ?sqd) / (?aMb - ?sqd)
     ?e2d = (?aMb - ?sqd) / (?aMb + ?sqd)"
    by (rule exp_2dist_formula)
  hence "cosh_dist ?pa ?pb
    = (sqrt ((?aMb + ?sqd) / (?aMb - ?sqd))
    + sqrt ((?aMb - ?sqd) / (?aMb + ?sqd)))
    / 2"
    by (unfold cosh_dist_def, auto)

  have "?qd  0"
  proof cases
    assume "?pa = ?pb"
    thus "?qd  0" by (simp add: quarter_discrim_self_zero)
  next
    assume "?pa  ?pb"
    with a  0 and b  0 and ?pa  hyp2
    have "?qd > 0" by (simp add: quarter_discrim_positive)
    thus "?qd  0" by simp
  qed
  with real_sqrt_pow2 [of ?qd] have "?sqd2 = ?qd" by simp
  hence "(?aMb + ?sqd) * (?aMb - ?sqd) = ?aMa * ?bMb"
    by (unfold quarter_discrim_def, simp add: algebra_simps power2_eq_square)

  from times_divide_times_eq [of
    "?aMb + ?sqd" "?aMb + ?sqd" "?aMb + ?sqd" "?aMb - ?sqd"]
  have "(?aMb + ?sqd) / (?aMb - ?sqd)
    = (?aMb + ?sqd)2 / ((?aMb + ?sqd) * (?aMb - ?sqd))"
    by (simp add: power2_eq_square)
  with (?aMb + ?sqd) * (?aMb - ?sqd) = ?aMa * ?bMb
  have "(?aMb + ?sqd) / (?aMb - ?sqd) = (?aMb + ?sqd)2 / (?aMa * ?bMb)" by simp
  hence "sqrt ((?aMb + ?sqd) / (?aMb - ?sqd))
    = ¦?aMb + ?sqd¦ / sqrt (?aMa * ?bMb)"
    by (simp add: real_sqrt_divide)

  from times_divide_times_eq [of
    "?aMb + ?sqd" "?aMb - ?sqd" "?aMb - ?sqd" "?aMb - ?sqd"]
  have "(?aMb - ?sqd) / (?aMb + ?sqd)
    = (?aMb - ?sqd)2 / ((?aMb + ?sqd) * (?aMb - ?sqd))"
    by (simp add: power2_eq_square)
  with (?aMb + ?sqd) * (?aMb - ?sqd) = ?aMa * ?bMb
  have "(?aMb - ?sqd) / (?aMb + ?sqd) = (?aMb - ?sqd)2 / (?aMa * ?bMb)" by simp
  hence "sqrt ((?aMb - ?sqd) / (?aMb + ?sqd))
    = ¦?aMb - ?sqd¦ / sqrt (?aMa * ?bMb)"
    by (simp add: real_sqrt_divide)

  from a  0 and b  0 and ?pa  hyp2 and ?pb  hyp2
  have "?aMa < 0" and "?bMb < 0"
    by (simp_all add: K2_imp_M_neg)
  with (?aMb + ?sqd) * (?aMb - ?sqd) = ?aMa * ?bMb
  have "(?aMb + ?sqd) * (?aMb - ?sqd) > 0" by (simp add: mult_neg_neg)
  hence "?aMb + ?sqd  0" and "?aMb - ?sqd  0" by auto
  hence "sgn (?aMb + ?sqd)  {-1,1}" and "sgn (?aMb - ?sqd)  {-1,1}"
    by (simp_all add: sgn_real_def)

  from (?aMb + ?sqd) * (?aMb - ?sqd) > 0
  have "sgn ((?aMb + ?sqd) * (?aMb - ?sqd)) = 1" by simp
  hence "sgn (?aMb + ?sqd) * sgn (?aMb - ?sqd) = 1" by (simp add: sgn_mult)
  with sgn (?aMb + ?sqd)  {-1,1} and sgn (?aMb - ?sqd)  {-1,1}
  have "sgn (?aMb + ?sqd) = sgn (?aMb - ?sqd)" by auto
  with abs_plus [of "?aMb + ?sqd" "?aMb - ?sqd"]
  have "¦?aMb + ?sqd¦ + ¦?aMb - ?sqd¦ = 2 * ¦?aMb¦" by simp
  with sqrt ((?aMb + ?sqd) / (?aMb - ?sqd))
    = ¦?aMb + ?sqd¦ / sqrt (?aMa * ?bMb)
    and sqrt ((?aMb - ?sqd) / (?aMb + ?sqd))
    = ¦?aMb - ?sqd¦ / sqrt (?aMa * ?bMb)
    and add_divide_distrib [of
    "¦?aMb + ?sqd¦" "¦?aMb - ?sqd¦" "sqrt (?aMa * ?bMb)"]
  have "sqrt ((?aMb + ?sqd) / (?aMb - ?sqd))
    + sqrt ((?aMb - ?sqd) / (?aMb + ?sqd))
    = 2 * ¦?aMb¦ / sqrt (?aMa * ?bMb)"
    by simp
  with cosh_dist ?pa ?pb
    = (sqrt ((?aMb + ?sqd) / (?aMb - ?sqd))
    + sqrt ((?aMb - ?sqd) / (?aMb + ?sqd)))
    / 2
  show "cosh_dist ?pa ?pb = ¦?aMb¦ / sqrt (?aMa * ?bMb)" by simp
qed

lemma cosh_dist_perp_special_case:
  assumes "¦x¦ < 1" and "¦y¦ < 1"
  shows "cosh_dist (proj2_abs (vector [x,0,1])) (proj2_abs (vector [0,y,1]))
  = (cosh_dist K2_centre (proj2_abs (vector [x,0,1])))
  * (cosh_dist K2_centre (proj2_abs (vector [0,y,1])))"
  (is "cosh_dist ?pa ?pb = (cosh_dist ?po ?pa) * (cosh_dist ?po ?pb)")
proof -
  have "vector [x,0,1]  (0::real^3)" (is "?a  0")
    and "vector [0,y,1]  (0::real^3)" (is "?b  0")
    by (unfold vector_def, simp_all add: vec_eq_iff forall_3)

  have "?a  (M *v ?a) = x2 - 1" (is "?aMa = x2 - 1")
    and "?b  (M *v ?b) = y2 - 1" (is "?bMb = y2 - 1")
    unfolding vector_def and M_def and inner_vec_def
      and matrix_vector_mult_def
    by (simp_all add: sum_3 power2_eq_square)
  with ¦x¦ < 1 and ¦y¦ < 1
  have "?aMa < 0" and "?bMb < 0" by (simp_all add: abs_square_less_1)
  hence "?pa  hyp2" and "?pb  hyp2"
    by (simp_all add: M_neg_imp_K2)
  with ?a  0 and ?b  0
  have "cosh_dist ?pa ?pb = ¦?a  (M *v ?b)¦ / sqrt (?aMa * ?bMb)"
    (is "cosh_dist ?pa ?pb = ¦?aMb¦ / sqrt (?aMa * ?bMb)")
    by (rule cosh_dist_formula)
  also from ?aMa = x2 - 1 and ?bMb = y2 - 1
  have " = ¦?aMb¦ / sqrt ((x2 - 1) * (y2 - 1))" by simp
  finally have "cosh_dist ?pa ?pb = 1 / sqrt ((1 - x2) * (1 - y2))"
    unfolding vector_def and M_def and inner_vec_def
      and matrix_vector_mult_def
    by (simp add: sum_3 algebra_simps)

  let ?o = "vector [0,0,1]"
  let ?oMa = "?o  (M *v ?a)"
  let ?oMb = "?o  (M *v ?b)"
  let ?oMo = "?o  (M *v ?o)"
  from K2_centre_non_zero and ?a  0 and ?b  0
    and K2_centre_in_K2 and ?pa  hyp2 and ?pb  hyp2
    and cosh_dist_formula [of ?o]
  have "cosh_dist ?po ?pa = ¦?oMa¦ / sqrt (?oMo * ?aMa)"
    and "cosh_dist ?po ?pb = ¦?oMb¦ / sqrt (?oMo * ?bMb)"
    by (unfold K2_centre_def, simp_all)
  hence "cosh_dist ?po ?pa = 1 / sqrt (1 - x2)"
    and "cosh_dist ?po ?pb = 1 / sqrt (1 - y2)"
    unfolding vector_def and M_def and inner_vec_def
      and matrix_vector_mult_def
    by (simp_all add: sum_3 power2_eq_square)
  with cosh_dist ?pa ?pb = 1 / sqrt ((1 - x2) * (1 - y2))
  show "cosh_dist ?pa ?pb = cosh_dist ?po ?pa * cosh_dist ?po ?pb"
    by (simp add: real_sqrt_mult)
qed

lemma K2_isometry_cross_ratio_endpoints_in_S:
  assumes "a  hyp2" and "b  hyp2" and "is_K2_isometry J" and "a  b"
  shows "cross_ratio (apply_cltn2 (endpoint_in_S a b) J)
  (apply_cltn2 (endpoint_in_S b a) J) (apply_cltn2 a J) (apply_cltn2 b J)
  = cross_ratio (endpoint_in_S a b) (endpoint_in_S b a) a b"
  (is "cross_ratio ?pJ ?qJ ?aJ ?bJ = cross_ratio ?p ?q a b")
proof -
  let ?l = "proj2_line_through a b"
  have "proj2_incident a ?l" and "proj2_incident b ?l"
    by (rule proj2_line_through_incident)+
  with a  b and a  hyp2 and b  hyp2
  have "proj2_incident ?p ?l" and "proj2_incident ?q ?l"
    by (simp_all add: endpoint_in_S_incident)
  with proj2_incident a ?l and proj2_incident b ?l
  have "proj2_set_Col {?p,?q,a,b}"
    by (unfold proj2_set_Col_def) (simp add: exI [of _ ?l])

  from a  b and a  hyp2 and b  hyp2
  have "?p  ?q" by (simp add: endpoint_in_S_swap)

  from a  hyp2 and b  hyp2 have "?p  S" by (simp add: endpoint_in_S)
  with a  hyp2 and b  hyp2
  have "a  ?p" and "b  ?p" by (simp_all add: hyp2_S_not_equal)
  with proj2_set_Col {?p,?q,a,b} and ?p  ?q
  show "cross_ratio ?pJ ?qJ ?aJ ?bJ = cross_ratio ?p ?q a b"
    by (rule cross_ratio_cltn2)
qed

lemma K2_isometry_exp_2dist:
  assumes "a  hyp2" and "b  hyp2" and "is_K2_isometry J"
  shows "exp_2dist (apply_cltn2 a J) (apply_cltn2 b J) = exp_2dist a b"
  (is "exp_2dist ?aJ ?bJ = _")
proof cases
  assume "a = b"
  thus "exp_2dist ?aJ ?bJ = exp_2dist a b" by (unfold exp_2dist_def) simp
next
  assume "a  b"
  with apply_cltn2_injective have "?aJ  ?bJ" by fast

  let ?p = "endpoint_in_S a b"
  let ?q = "endpoint_in_S b a"
  let ?aJ = "apply_cltn2 a J"
    and ?bJ = "apply_cltn2 b J"
    and ?pJ = "apply_cltn2 ?p J"
    and ?qJ = "apply_cltn2 ?q J"
  from a  b and a  hyp2 and b  hyp2 and is_K2_isometry J
  have "endpoint_in_S ?aJ ?bJ = ?pJ" and "endpoint_in_S ?bJ ?aJ = ?qJ"
    by (simp_all add: K2_isometry_endpoint_in_S)

  from assms and a  b
  have "cross_ratio ?pJ ?qJ ?aJ ?bJ = cross_ratio ?p ?q a b"
    by (rule K2_isometry_cross_ratio_endpoints_in_S)
  with endpoint_in_S ?aJ ?bJ = ?pJ and endpoint_in_S ?bJ ?aJ = ?qJ
    and a  b and ?aJ  ?bJ
  show "exp_2dist ?aJ ?bJ = exp_2dist a b" by (unfold exp_2dist_def) simp
qed

lemma K2_isometry_cosh_dist:
  assumes "a  hyp2" and "b  hyp2" and "is_K2_isometry J"
  shows "cosh_dist (apply_cltn2 a J) (apply_cltn2 b J) = cosh_dist a b"
  using assms
  by (unfold cosh_dist_def) (simp add: K2_isometry_exp_2dist)

lemma cosh_dist_perp:
  assumes "M_perp l m" and "a  hyp2" and "b  hyp2" and "c  hyp2"
  and "proj2_incident a l" and "proj2_incident b l"
  and "proj2_incident b m" and "proj2_incident c m"
  shows "cosh_dist a c = cosh_dist b a * cosh_dist b c"
proof -
  from M_perp l m and b  hyp2 and proj2_incident b l
    and proj2_incident b m and M_perp_to_compass [of l m b b]
  obtain J where "is_K2_isometry J" and "apply_cltn2_line equator J = l"
    and "apply_cltn2_line meridian J = m"
    by auto

  let ?Ji = "cltn2_inverse J"
  let ?aJi = "apply_cltn2 a ?Ji"
  let ?bJi = "apply_cltn2 b ?Ji"
  let ?cJi = "apply_cltn2 c ?Ji"
  from apply_cltn2_line equator J = l and apply_cltn2_line meridian J = m
    and proj2_incident a l and proj2_incident b l
    and proj2_incident b m and proj2_incident c m
  have "proj2_incident ?aJi equator" and "proj2_incident ?bJi equator"
    and "proj2_incident ?bJi meridian" and "proj2_incident ?cJi meridian"
    by (auto simp add: apply_cltn2_incident)

  from is_K2_isometry J
  have "is_K2_isometry ?Ji" by (rule cltn2_inverse_is_K2_isometry)
  with a  hyp2 and c  hyp2
  have "?aJi  hyp2" and "?cJi  hyp2"
    by (simp_all add: statement60_one_way)

  from ?aJi  hyp2 and proj2_incident ?aJi equator
    and on_equator_in_hyp2_rep
  obtain x where "¦x¦ < 1" and "?aJi = proj2_abs (vector [x,0,1])" by auto
  moreover
  from ?cJi  hyp2 and proj2_incident ?cJi meridian
    and on_meridian_in_hyp2_rep
  obtain y where "¦y¦ < 1" and "?cJi = proj2_abs (vector [0,y,1])" by auto
  moreover
  from proj2_incident ?bJi equator and proj2_incident ?bJi meridian
  have "?bJi = K2_centre" by (rule on_equator_meridian_is_K2_centre)
  ultimately
  have "cosh_dist ?aJi ?cJi = cosh_dist ?bJi ?aJi * cosh_dist ?bJi ?cJi"
    by (simp add: cosh_dist_perp_special_case)
  with a  hyp2 and b  hyp2 and c  hyp2 and is_K2_isometry ?Ji
  show "cosh_dist a c = cosh_dist b a * cosh_dist b c"
    by (simp add: K2_isometry_cosh_dist)
qed

lemma are_endpoints_in_S_ordered_cross_ratio:
  assumes "are_endpoints_in_S p q a b"
  and "B (cart2_pt a) (cart2_pt b) (cart2_pt p)" (is "B ?ca ?cb ?cp")
  shows "cross_ratio p q a b  1"
proof -
  from are_endpoints_in_S p q a b
  have "p  q" and "p  S" and "q  S" and "a  hyp2" and "b  hyp2"
    and "proj2_set_Col {p,q,a,b}"
     by (unfold are_endpoints_in_S_def) simp_all

   from a  hyp2 and b  hyp2 and p  S and q  S
   have "z_non_zero a" and "z_non_zero b" and "z_non_zero p" and "z_non_zero q"
     by (simp_all add: hyp2_S_z_non_zero)
  hence "proj2_abs (cart2_append1 p) = p" (is "proj2_abs ?cp1 = p")
    and "proj2_abs (cart2_append1 q) = q" (is "proj2_abs ?cq1 = q")
    and "proj2_abs (cart2_append1 a) = a" (is "proj2_abs ?ca1 = a")
    and "proj2_abs (cart2_append1 b) = b" (is "proj2_abs ?cb1 = b")
    by (simp_all add: proj2_abs_cart2_append1)

   from b  hyp2 and p  S have "b  p" by (rule hyp2_S_not_equal)
   with z_non_zero a and z_non_zero b and z_non_zero p
     and B ?ca ?cb ?cp and cart2_append1_between_right_strict [of a b p]
   obtain j where "j  0" and "j < 1" and "?cb1 = j *R ?cp1 + (1-j) *R ?ca1"
     by auto

   from proj2_set_Col {p,q,a,b}
   obtain l where "proj2_incident q l" and "proj2_incident p l"
     and "proj2_incident a l"
     by (unfold proj2_set_Col_def) auto
   with p  q and q  S and p  S and a  hyp2
     and S_hyp2_S_cart2_append1 [of q p a l]
   obtain k where "k > 0" and "k < 1" and "?ca1 = k *R ?cp1 + (1-k) *R ?cq1"
     by auto

   from z_non_zero p and z_non_zero q
   have "?cp1  0" and "?cq1  0" by (simp_all add: cart2_append1_non_zero)

   from p  q and proj2_abs ?cp1 = p and proj2_abs ?cq1 = q
   have "proj2_abs ?cp1  proj2_abs ?cq1" by simp

   from k < 1 have "1-k  0" by simp
   with j < 1 have "(1-j)*(1-k)  0" by simp

   from j < 1 and k > 0 have "(1-j)*k > 0" by simp

   from ?cb1 = j *R ?cp1 + (1-j) *R ?ca1
   have "?cb1 = (j+(1-j)*k) *R ?cp1 + ((1-j)*(1-k)) *R ?cq1"
     by (unfold ?ca1 = k *R ?cp1 + (1-k) *R ?cq1) (simp add: algebra_simps)
   with ?ca1 = k *R ?cp1 + (1-k) *R ?cq1
   have "proj2_abs ?ca1 = proj2_abs (k *R ?cp1 + (1-k) *R ?cq1)"
     and "proj2_abs ?cb1
     = proj2_abs ((j+(1-j)*k) *R ?cp1 + ((1-j)*(1-k)) *R ?cq1)"
     by simp_all
   with proj2_abs ?ca1 = a and proj2_abs ?cb1 = b
   have "a = proj2_abs (k *R ?cp1 + (1-k) *R ?cq1)"
     and "b = proj2_abs ((j+(1-j)*k) *R ?cp1 + ((1-j)*(1-k)) *R ?cq1)"
     by simp_all
   with proj2_abs ?cp1 = p and proj2_abs ?cq1 = q
   have "cross_ratio p q a b
     = cross_ratio (proj2_abs ?cp1) (proj2_abs ?cq1)
     (proj2_abs (k *R ?cp1 + (1-k) *R ?cq1))
     (proj2_abs ((j+(1-j)*k) *R ?cp1 + ((1-j)*(1-k)) *R ?cq1))"
     by simp
   also from ?cp1  0 and ?cq1  0 and proj2_abs ?cp1  proj2_abs ?cq1
     and 1-k  0 and (1-j)*(1-k)  0
   have " = (1-k)*(j+(1-j)*k) / (k*((1-j)*(1-k)))" by (rule cross_ratio_abs)
   also from 1-k  0 have " = (j+(1-j)*k) / ((1-j)*k)" by simp
   also from j  0 and (1-j)*k > 0 have "  1" by simp
   finally show "cross_ratio p q a b  1" .
qed

lemma cross_ratio_S_S_hyp2_hyp2_positive:
  assumes "are_endpoints_in_S p q a b"
  shows "cross_ratio p q a b > 0"
proof cases
  assume "B (cart2_pt p) (cart2_pt b) (cart2_pt a)"
  hence "B (cart2_pt a) (cart2_pt b) (cart2_pt p)"
    by (rule real_euclid.th3_2)
  with assms have "cross_ratio p q a b  1"
    by (rule are_endpoints_in_S_ordered_cross_ratio)
  thus "cross_ratio p q a b > 0" by simp
next
  assume "¬ B (cart2_pt p) (cart2_pt b) (cart2_pt a)" (is "¬ B ?cp ?cb ?ca")

  from are_endpoints_in_S p q a b
  have "are_endpoints_in_S p q b a" by (rule are_endpoints_in_S_swap_34)

  from are_endpoints_in_S p q a b
  have "p  S" and "a  hyp2" and "b  hyp2" and "proj2_set_Col {p,q,a,b}"
    by (unfold are_endpoints_in_S_def) simp_all

  from proj2_set_Col {p,q,a,b}
  have "proj2_set_Col {p,a,b}"
    by (simp add: proj2_subset_Col [of "{p,a,b}" "{p,q,a,b}"])
  hence "proj2_Col p a b" by (subst proj2_Col_iff_set_Col)
  with p  S and a  hyp2 and b  hyp2
  have "B ?cp ?ca ?cb  B ?cp ?cb ?ca" by (simp add: S_at_edge)
  with ¬ B ?cp ?cb ?ca have "B ?cp ?ca ?cb" by simp
  hence "B ?cb ?ca ?cp" by (rule real_euclid.th3_2)
  with are_endpoints_in_S p q b a
  have "cross_ratio p q b a  1"
    by (rule are_endpoints_in_S_ordered_cross_ratio)
  thus "cross_ratio p q a b > 0" by (subst cross_ratio_swap_34) simp
qed

lemma cosh_dist_general:
  assumes "are_endpoints_in_S p q a b"
  shows "cosh_dist a b
  = (sqrt (cross_ratio p q a b) + 1 / sqrt (cross_ratio p q a b)) / 2"
proof -
  from are_endpoints_in_S p q a b
  have "p  q" and "p  S" and "q  S" and "a  hyp2" and "b  hyp2"
    and "proj2_set_Col {p,q,a,b}"
    by (unfold are_endpoints_in_S_def) simp_all

  from a  hyp2 and b  hyp2 and p  S and q  S
  have "a  p" and "a  q" and "b  p" and "b  q"
    by (simp_all add: hyp2_S_not_equal)

  show "cosh_dist a b
    = (sqrt (cross_ratio p q a b) + 1 / sqrt (cross_ratio p q a b)) / 2"
  proof cases
    assume "a = b"
    hence "cosh_dist a b = 1" by (unfold cosh_dist_def exp_2dist_def) simp

    from proj2_set_Col {p,q,a,b}
    have "proj2_Col p q a" by (unfold a = b) (simp add: proj2_Col_iff_set_Col)
    with p  q and a  p and a  q
    have "cross_ratio p q a b = 1" by (simp add: a = b cross_ratio_equal_1)
    hence "(sqrt (cross_ratio p q a b) + 1 / sqrt (cross_ratio p q a b)) / 2
      = 1"
      by simp
    with cosh_dist a b = 1
    show "cosh_dist a b
      = (sqrt (cross_ratio p q a b) + 1 / sqrt (cross_ratio p q a b)) / 2"
      by simp
  next
    assume "a  b"

    let ?r = "endpoint_in_S a b"
    let ?s = "endpoint_in_S b a"
    from a  b
    have "exp_2dist a b = cross_ratio ?r ?s a b" by (unfold exp_2dist_def) simp

    from a  b and are_endpoints_in_S p q a b
    have "(p = ?r  q = ?s)  (q = ?r  p = ?s)" by (rule are_endpoints_in_S)

    show "cosh_dist a b
      = (sqrt (cross_ratio p q a b) + 1 / sqrt (cross_ratio p q a b)) / 2"
    proof cases
      assume "p = ?r  q = ?s"
      with exp_2dist a b = cross_ratio ?r ?s a b
      have "exp_2dist a b = cross_ratio p q a b" by simp
      thus "cosh_dist a b
        = (sqrt (cross_ratio p q a b) + 1 / sqrt (cross_ratio p q a b)) / 2"
        by (unfold cosh_dist_def) (simp add: real_sqrt_divide)
    next
      assume "¬ (p = ?r  q = ?s)"
      with (p = ?r  q = ?s)  (q = ?r  p = ?s)
      have "q = ?r" and "p = ?s" by simp_all
      with exp_2dist a b = cross_ratio ?r ?s a b
      have "exp_2dist a b = cross_ratio q p a b" by simp

      have "{q,p,a,b} = {p,q,a,b}" by auto
      with proj2_set_Col {p,q,a,b} and p  q and a  p and b  p
        and a  q and b  q
      have "cross_ratio_correct p q a b" and "cross_ratio_correct q p a b"
        by (unfold cross_ratio_correct_def) simp_all
      hence "cross_ratio q p a b = 1 / (cross_ratio p q a b)"
        by (rule cross_ratio_swap_12)
      with exp_2dist a b = cross_ratio q p a b
      have "exp_2dist a b = 1 / (cross_ratio p q a b)" by simp
      thus "cosh_dist a b
        = (sqrt (cross_ratio p q a b) + 1 / sqrt (cross_ratio p q a b)) / 2"
        by (unfold cosh_dist_def) (simp add: real_sqrt_divide)
    qed
  qed
qed

lemma exp_2dist_positive:
  assumes "a  hyp2" and "b  hyp2"
  shows "exp_2dist a b > 0"
proof cases
  assume "a = b"
  thus "exp_2dist a b > 0" by (unfold exp_2dist_def) simp
next
  assume "a  b"

  let ?p = "endpoint_in_S a b"
  let ?q = "endpoint_in_S b a"
  from a  b and a  hyp2 and b  hyp2
  have "are_endpoints_in_S ?p ?q a b"
    by (rule endpoints_in_S_are_endpoints_in_S)
  hence "cross_ratio ?p ?q a b > 0" by (rule cross_ratio_S_S_hyp2_hyp2_positive)
  with a  b show "exp_2dist a b > 0" by (unfold exp_2dist_def) simp
qed

lemma cosh_dist_at_least_1:
  assumes "a  hyp2" and "b  hyp2"
  shows "cosh_dist a b  1"
proof -
  from assms have "exp_2dist a b > 0" by (rule exp_2dist_positive)
  with am_gm2(1) [of "sqrt (exp_2dist a b)" "sqrt (1 / exp_2dist a b)"]
  show "cosh_dist a b  1"
    by (unfold cosh_dist_def) (simp add: real_sqrt_mult [symmetric])
qed

lemma cosh_dist_positive:
  assumes "a  hyp2" and "b  hyp2"
  shows "cosh_dist a b > 0"
proof -
  from assms have "cosh_dist a b  1" by (rule cosh_dist_at_least_1)
  thus "cosh_dist a b > 0" by simp
qed

lemma cosh_dist_perp_divide:
  assumes "M_perp l m" and "a  hyp2" and "b  hyp2" and "c  hyp2"
  and "proj2_incident a l" and "proj2_incident b l" and "proj2_incident b m"
  and "proj2_incident c m"
  shows "cosh_dist b c = cosh_dist a c / cosh_dist b a"
proof -
  from b  hyp2 and a  hyp2
  have "cosh_dist b a > 0" by (rule cosh_dist_positive)

  from assms
  have "cosh_dist a c = cosh_dist b a * cosh_dist b c" by (rule cosh_dist_perp)
  with cosh_dist b a > 0
  show "cosh_dist b c = cosh_dist a c / cosh_dist b a" by simp
qed

lemma real_hyp2_C_cross_ratio_endpoints_in_S:
  assumes "a  b" and "a b \<congruent>K c d"
  shows "cross_ratio (endpoint_in_S (Rep_hyp2 a) (Rep_hyp2 b))
  (endpoint_in_S (Rep_hyp2 b) (Rep_hyp2 a)) (Rep_hyp2 a) (Rep_hyp2 b)
  = cross_ratio (endpoint_in_S (Rep_hyp2 c) (Rep_hyp2 d))
  (endpoint_in_S (Rep_hyp2 d) (Rep_hyp2 c)) (Rep_hyp2 c) (Rep_hyp2 d)"
  (is "cross_ratio ?p ?q ?a' ?b' = cross_ratio ?r ?s ?c' ?d'")
proof -
  from a  b and a b \<congruent>K c d have "c  d" by (auto simp add: hyp2.A3')
  with a  b have "?a'  ?b'" and "?c'  ?d'" by (unfold Rep_hyp2_inject)

  from a b \<congruent>K c d
  obtain J where "is_K2_isometry J" and "hyp2_cltn2 a J = c"
    and "hyp2_cltn2 b J = d"
    by (unfold real_hyp2_C_def) auto
  hence "apply_cltn2 ?a' J = ?c'" and "apply_cltn2 ?b' J = ?d'"
    by (simp_all add: Rep_hyp2_cltn2 [symmetric])
  with ?a'  ?b' and is_K2_isometry J
  have "apply_cltn2 ?p J = ?r" and "apply_cltn2 ?q J = ?s"
    by (simp_all add: Rep_hyp2 K2_isometry_endpoint_in_S)

  from ?a'  ?b'
  have "proj2_set_Col {?p,?q,?a',?b'}"
    by (simp add: Rep_hyp2 proj2_set_Col_endpoints_in_S)

  from ?a'  ?b' have "?p  ?q" by (simp add: Rep_hyp2 endpoint_in_S_swap)

  have "?p  S" by (simp add: Rep_hyp2 endpoint_in_S)
  hence "?a'  ?p" and "?b'  ?p" by (simp_all add: Rep_hyp2 hyp2_S_not_equal)
  with proj2_set_Col {?p,?q,?a',?b'} and ?p  ?q
  have "cross_ratio ?p ?q ?a' ?b'
    = cross_ratio (apply_cltn2 ?p J) (apply_cltn2 ?q J)
    (apply_cltn2 ?a' J) (apply_cltn2 ?b' J)"
    by (rule cross_ratio_cltn2 [symmetric])
  with apply_cltn2 ?p J = ?r and apply_cltn2 ?q J = ?s
    and apply_cltn2 ?a' J = ?c' and apply_cltn2 ?b' J = ?d'
  show "cross_ratio ?p ?q ?a' ?b' = cross_ratio ?r ?s ?c' ?d'" by simp
qed

lemma real_hyp2_C_exp_2dist:
  assumes "a b \<congruent>K c d"
  shows "exp_2dist (Rep_hyp2 a) (Rep_hyp2 b)
  = exp_2dist (Rep_hyp2 c) (Rep_hyp2 d)"
  (is "exp_2dist ?a' ?b' = exp_2dist ?c' ?d'")
proof -
  from a b \<congruent>K c d
  obtain J where "is_K2_isometry J" and "hyp2_cltn2 a J = c"
    and "hyp2_cltn2 b J = d"
    by (unfold real_hyp2_C_def) auto
  hence "apply_cltn2 ?a' J = ?c'" and "apply_cltn2 ?b' J = ?d'"
    by (simp_all add: Rep_hyp2_cltn2 [symmetric])

  from Rep_hyp2 [of a] and Rep_hyp2 [of b] and is_K2_isometry J
  have "exp_2dist (apply_cltn2 ?a' J) (apply_cltn2 ?b' J) = exp_2dist ?a' ?b'"
    by (rule K2_isometry_exp_2dist)
  with apply_cltn2 ?a' J = ?c' and apply_cltn2 ?b' J = ?d'
  show "exp_2dist ?a' ?b' = exp_2dist ?c' ?d'" by simp
qed

lemma real_hyp2_C_cosh_dist:
  assumes "a b \<congruent>K c d"
  shows "cosh_dist (Rep_hyp2 a) (Rep_hyp2 b)
  = cosh_dist (Rep_hyp2 c) (Rep_hyp2 d)"
  using assms
  by (unfold cosh_dist_def) (simp add: real_hyp2_C_exp_2dist)

lemma cross_ratio_in_terms_of_cosh_dist:
  assumes "are_endpoints_in_S p q a b"
  and "B (cart2_pt a) (cart2_pt b) (cart2_pt p)"
  shows "cross_ratio p q a b
  = 2 * (cosh_dist a b)2 + 2 * cosh_dist a b * sqrt ((cosh_dist a b)2 - 1) - 1"
  (is "?pqab = 2 * ?ab2 + 2 * ?ab * sqrt (?ab2 - 1) - 1")
proof -
  from are_endpoints_in_S p q a b
  have "?ab = (sqrt ?pqab + 1 / sqrt ?pqab) / 2" by (rule cosh_dist_general)
  hence "sqrt ?pqab - 2 * ?ab + 1 / sqrt ?pqab = 0" by simp
  hence "sqrt ?pqab * (sqrt ?pqab - 2 * ?ab + 1 / sqrt ?pqab) = 0" by simp
  moreover from assms
  have "?pqab  1" by (rule are_endpoints_in_S_ordered_cross_ratio)
  ultimately have "?pqab - 2 * ?ab * (sqrt ?pqab) + 1 = 0"
    by (simp add: algebra_simps real_sqrt_mult [symmetric])
  with ?pqab  1 and discriminant_iff [of 1 "sqrt ?pqab" "- 2 * ?ab" 1]
  have "sqrt ?pqab = (2 * ?ab + sqrt (4 * ?ab2 - 4)) / 2
     sqrt ?pqab = (2 * ?ab - sqrt (4 * ?ab2 - 4)) / 2"
    unfolding discrim_def
    by (simp add: real_sqrt_mult [symmetric] power2_eq_square)
  moreover have "sqrt (4 * ?ab2 - 4) = sqrt (4 * (?ab2 - 1))" by simp
  hence "sqrt (4 * ?ab2 - 4) = 2 * sqrt (?ab2 - 1)"
    by (unfold real_sqrt_mult) simp
  ultimately have "sqrt ?pqab = 2 * (?ab + sqrt (?ab2 - 1)) / 2
     sqrt ?pqab = 2 * (?ab - sqrt (?ab2 - 1)) / 2"
    by simp
  hence "sqrt ?pqab = ?ab + sqrt (?ab2 - 1)
     sqrt ?pqab = ?ab - sqrt (?ab2 - 1)"
    by (simp only: nonzero_mult_div_cancel_left [of 2])

  from are_endpoints_in_S p q a b
  have "a  hyp2" and "b  hyp2" by (unfold are_endpoints_in_S_def) simp_all
  hence "?ab  1" by (rule cosh_dist_at_least_1)
  hence "?ab2  1" by simp
  hence "sqrt (?ab2 - 1)  0" by simp
  hence "sqrt (?ab2 - 1) * sqrt (?ab2 - 1) = ?ab2 - 1"
    by (simp add: real_sqrt_mult [symmetric])
  hence "(?ab + sqrt (?ab2 - 1)) * (?ab - sqrt (?ab2 - 1)) = 1"
    by (simp add: algebra_simps power2_eq_square)

  have "?ab - sqrt (?ab2 - 1)  1"
  proof (rule ccontr)
    assume "¬ (?ab - sqrt (?ab2 - 1)  1)"
    hence "1 < ?ab - sqrt (?ab2 - 1)" by simp
    also from sqrt (?ab2 - 1)  0
    have "  ?ab + sqrt (?ab2 - 1)" by simp
    finally have "1 < ?ab + sqrt (?ab2 - 1)" by simp
    with 1 < ?ab - sqrt (?ab2 - 1)
      and mult_strict_mono' [of
      1 "?ab + sqrt (?ab2 - 1)" 1 "?ab - sqrt (?ab2 - 1)"]
    have "1 < (?ab + sqrt (?ab2 - 1)) * (?ab - sqrt (?ab2 - 1))" by simp
    with (?ab + sqrt (?ab2 - 1)) * (?ab - sqrt (?ab2 - 1)) = 1
    show False by simp
  qed

  have "sqrt ?pqab = ?ab + sqrt (?ab2 - 1)"
  proof (rule ccontr)
    assume "sqrt ?pqab  ?ab + sqrt (?ab2 - 1)"
    with sqrt ?pqab = ?ab + sqrt (?ab2 - 1)
       sqrt ?pqab = ?ab - sqrt (?ab2 - 1)
    have "sqrt ?pqab = ?ab - sqrt (?ab2 - 1)" by simp
    with ?ab - sqrt (?ab2 - 1)  1 have "sqrt ?pqab  1" by simp
    with ?pqab  1 have "sqrt ?pqab = 1" by simp
    with sqrt ?pqab = ?ab - sqrt (?ab2 - 1)
      and (?ab + sqrt (?ab2 - 1)) * (?ab - sqrt (?ab2 - 1)) = 1
    have "?ab + sqrt (?ab2 - 1) = 1" by simp
    with sqrt ?pqab = 1 have "sqrt ?pqab = ?ab + sqrt (?ab2 - 1)" by simp
    with sqrt ?pqab  ?ab + sqrt (?ab2 - 1) show False ..
  qed
  moreover from ?pqab  1 have "?pqab = (sqrt ?pqab)2" by simp
  ultimately have "?pqab = (?ab + sqrt (?ab2 - 1))2" by simp
  with sqrt (?ab2 - 1) * sqrt (?ab2 - 1) = ?ab2 - 1
  show "?pqab = 2 * ?ab2 + 2 * ?ab * sqrt (?ab2 - 1) - 1"
    by (simp add: power2_eq_square algebra_simps)
qed

lemma are_endpoints_in_S_cross_ratio_correct:
  assumes "are_endpoints_in_S p q a b"
  shows "cross_ratio_correct p q a b"
proof -
  from are_endpoints_in_S p q a b
  have "p  q" and "p  S" and "q  S" and "a  hyp2" and "b  hyp2"
    and "proj2_set_Col {p,q,a,b}"
    by (unfold are_endpoints_in_S_def) simp_all

  from a  hyp2 and b  hyp2 and p  S and q  S
  have "a  p" and "b  p" and "a  q" by (simp_all add: hyp2_S_not_equal)
  with proj2_set_Col {p,q,a,b} and p  q
  show "cross_ratio_correct p q a b" by (unfold cross_ratio_correct_def) simp
qed

lemma endpoints_in_S_cross_ratio_correct:
  assumes "a  b" and "a  hyp2" and "b  hyp2"
  shows "cross_ratio_correct (endpoint_in_S a b) (endpoint_in_S b a) a b"
proof -
  from assms
  have "are_endpoints_in_S (endpoint_in_S a b) (endpoint_in_S b a) a b"
    by (rule endpoints_in_S_are_endpoints_in_S)
  thus "cross_ratio_correct (endpoint_in_S a b) (endpoint_in_S b a) a b"
    by (rule are_endpoints_in_S_cross_ratio_correct)
qed

lemma endpoints_in_S_perp_foot_cross_ratio_correct:
  assumes "a  hyp2" and "b  hyp2" and "c  hyp2" and "a  b"
  and "proj2_incident a l" and "proj2_incident b l"
  shows "cross_ratio_correct
  (endpoint_in_S a b) (endpoint_in_S b a) a (perp_foot c l)"
  (is "cross_ratio_correct ?p ?q a ?d")
proof -
  from assms
  have "are_endpoints_in_S ?p ?q a ?d"
    by (rule endpoints_in_S_perp_foot_are_endpoints_in_S)
  thus "cross_ratio_correct ?p ?q a ?d"
    by (rule are_endpoints_in_S_cross_ratio_correct)
qed

lemma cosh_dist_unique:
  assumes "a  hyp2" and "b  hyp2" and "c  hyp2" and "p  S"
  and "B (cart2_pt a) (cart2_pt b) (cart2_pt p)" (is "B ?ca ?cb ?cp")
  and "B (cart2_pt a) (cart2_pt c) (cart2_pt p)" (is "B ?ca ?cc ?cp")
  and "cosh_dist a b = cosh_dist a c" (is "?ab = ?ac")
  shows "b = c"
proof -
  let ?q = "endpoint_in_S p a"

  from a  hyp2 and b  hyp2 and c  hyp2 and p  S
  have "z_non_zero a" and "z_non_zero b" and "z_non_zero c" and "z_non_zero p"
    by (simp_all add: hyp2_S_z_non_zero)
  with B ?ca ?cb ?cp and B ?ca ?cc ?cp
  have " l. proj2_incident a l  proj2_incident b l  proj2_incident p l"
    and " m. proj2_incident a m  proj2_incident c m  proj2_incident p m"
    by (simp_all add: euclid_B_cart2_common_line)
  then obtain l and m where
    "proj2_incident a l" and "proj2_incident b l" and "proj2_incident p l"
    and "proj2_incident a m" and "proj2_incident c m" and "proj2_incident p m"
    by auto

  from a  hyp2 and p  S have "a  p" by (rule hyp2_S_not_equal)
  with proj2_incident a l and proj2_incident p l
    and proj2_incident a m and proj2_incident p m and proj2_incident_unique
  have "l = m" by fast
  with proj2_incident c m have "proj2_incident c l" by simp
  with a  hyp2 and b  hyp2 and c  hyp2 and p  S
    and proj2_incident a l and proj2_incident b l and proj2_incident p l
  have "are_endpoints_in_S p ?q b a" and "are_endpoints_in_S p ?q c a"
    by (simp_all add: end_and_opposite_are_endpoints_in_S)
  with are_endpoints_in_S_swap_34
  have "are_endpoints_in_S p ?q a b" and "are_endpoints_in_S p ?q a c" by fast+
  hence "cross_ratio_correct p ?q a b" and "cross_ratio_correct p ?q a c"
    by (simp_all add: are_endpoints_in_S_cross_ratio_correct)
  moreover
  from are_endpoints_in_S p ?q a b and are_endpoints_in_S p ?q a c
    and B ?ca ?cb ?cp and B ?ca ?cc ?cp
  have "cross_ratio p ?q a b = 2 * ?ab2 + 2 * ?ab * sqrt (?ab2 - 1) - 1"
    and "cross_ratio p ?q a c = 2 * ?ac2 + 2 * ?ac * sqrt (?ac2 - 1) - 1"
    by (simp_all add: cross_ratio_in_terms_of_cosh_dist)
  with ?ab = ?ac have "cross_ratio p ?q a b = cross_ratio p ?q a c" by simp
  ultimately show "b = c" by (rule cross_ratio_unique)
qed

lemma cosh_dist_swap:
  assumes "a  hyp2" and "b  hyp2"
  shows "cosh_dist a b = cosh_dist b a"
proof -
  from assms and K2_isometry_swap
  obtain J where "is_K2_isometry J" and "apply_cltn2 a J = b"
    and "apply_cltn2 b J = a"
    by auto

  from b  hyp2 and a  hyp2 and is_K2_isometry J
  have "cosh_dist (apply_cltn2 b J) (apply_cltn2 a J) = cosh_dist b a"
    by (rule K2_isometry_cosh_dist)
  with apply_cltn2 a J = b and apply_cltn2 b J = a
  show "cosh_dist a b = cosh_dist b a" by simp
qed

lemma exp_2dist_1_equal:
  assumes "a  hyp2" and "b  hyp2" and "exp_2dist a b = 1"
  shows "a = b"
proof (rule ccontr)
  assume "a  b"
  with a  hyp2 and b  hyp2
  have "cross_ratio_correct (endpoint_in_S a b) (endpoint_in_S b a) a b"
    (is "cross_ratio_correct ?p ?q a b")
    by (simp add: endpoints_in_S_cross_ratio_correct)
  moreover
  from a  b
  have "exp_2dist a b = cross_ratio ?p ?q a b" by (unfold exp_2dist_def) simp
  with exp_2dist a b = 1 have "cross_ratio ?p ?q a b = 1" by simp
  ultimately have "a = b" by (rule cross_ratio_1_equal)
  with a  b show False ..
qed

subsubsection ‹A formula for a cross ratio involving a perpendicular foot›

lemma described_perp_foot_cross_ratio_formula:
  assumes "a  b" and "c  hyp2" and "are_endpoints_in_S p q a b"
  and "proj2_incident p l" and "proj2_incident q l" and "M_perp l m"
  and "proj2_incident d l" and "proj2_incident d m" and "proj2_incident c m"
  shows "cross_ratio p q d a
    = (cosh_dist b c * sqrt (cross_ratio p q a b) - cosh_dist a c)
      / (cosh_dist a c * cross_ratio p q a b
        - cosh_dist b c * sqrt (cross_ratio p q a b))"
  (is "?pqda = (?bc * sqrt ?pqab - ?ac) / (?ac * ?pqab - ?bc * sqrt ?pqab)")
proof -
  let ?da = "cosh_dist d a"
  let ?db = "cosh_dist d b"
  let ?dc = "cosh_dist d c"
  let ?pqdb = "cross_ratio p q d b"

  from are_endpoints_in_S p q a b
  have "p  q" and "p  S" and "q  S" and "a  hyp2" and "b  hyp2"
    and "proj2_set_Col {p,q,a,b}"
    by (unfold are_endpoints_in_S_def) simp_all

  from proj2_set_Col {p,q,a,b}
  obtain l' where "proj2_incident p l'" and "proj2_incident q l'"
    and "proj2_incident a l'" and "proj2_incident b l'"
    by (unfold proj2_set_Col_def) auto

  from p  q and proj2_incident p l' and proj2_incident q l'
    and proj2_incident p l and proj2_incident q l and proj2_incident_unique
  have "l' = l" by fast
  with proj2_incident a l' and proj2_incident b l'
  have "proj2_incident a l" and "proj2_incident b l" by simp_all

  from M_perp l m and a  hyp2 and proj2_incident a l and c  hyp2
    and proj2_incident c m and proj2_incident d l and proj2_incident d m
  have "d  hyp2" by (rule M_perp_hyp2)
  with a  hyp2 and b  hyp2 and c  hyp2
  have "?bc > 0" and "?da > 0" and "?ac > 0"
    by (simp_all add: cosh_dist_positive)

  from proj2_incident p l and proj2_incident q l and proj2_incident d l
    and proj2_incident a l and proj2_incident b l
  have "proj2_set_Col {p,q,d,a}" and "proj2_set_Col {p,q,d,b}"
    and "proj2_set_Col {p,q,a,b}"
    by (unfold proj2_set_Col_def) (simp_all add: exI [of _ l])
  with p  q and p  S and q  S and d  hyp2 and a  hyp2
    and b  hyp2
  have "are_endpoints_in_S p q d a" and "are_endpoints_in_S p q d b"
    and "are_endpoints_in_S p q a b"
    by (unfold are_endpoints_in_S_def) simp_all
  hence "?pqda > 0" and "?pqdb > 0" and "?pqab > 0"
    by (simp_all add: cross_ratio_S_S_hyp2_hyp2_positive)

  from proj2_incident p l and proj2_incident q l and proj2_incident a l
  have "proj2_Col p q a" by (rule proj2_incident_Col)

  from a  hyp2 and b  hyp2 and p  S and q  S
  have "a  p" and "a  q" and "b  p" by (simp_all add: hyp2_S_not_equal)

  from proj2_Col p q a and p  q and a  p and a  q
  have "?pqdb = ?pqda * ?pqab" by (rule cross_ratio_product [symmetric])

  from M_perp l m and a  hyp2 and b  hyp2 and c  hyp2 and d  hyp2
    and proj2_incident a l and proj2_incident b l and proj2_incident d l
    and proj2_incident d m and proj2_incident c m
    and cosh_dist_perp_divide [of l m _ d c]
  have "?dc = ?ac / ?da" and "?dc = ?bc / ?db" by fast+
  hence "?ac / ?da = ?bc / ?db" by simp
  with ?bc > 0 and ?da > 0
  have "?ac / ?bc = ?da / ?db" by (simp add: field_simps)
  also from are_endpoints_in_S p q d a and are_endpoints_in_S p q d b
  have "
    = 2 * (sqrt ?pqda + 1 / (sqrt ?pqda))
    / (2 * (sqrt ?pqdb + 1 / (sqrt ?pqdb)))"
    by (simp add: cosh_dist_general)
  also
  have " = (sqrt ?pqda + 1 / (sqrt ?pqda)) / (sqrt ?pqdb + 1 / (sqrt ?pqdb))"
    by (simp only: mult_divide_mult_cancel_left_if) simp
  also have "
    = sqrt ?pqdb * (sqrt ?pqda + 1 / (sqrt ?pqda))
    / (sqrt ?pqdb * (sqrt ?pqdb + 1 / (sqrt ?pqdb)))"
    by simp
  also from ?pqdb > 0
  have " = (sqrt (?pqdb * ?pqda) + sqrt (?pqdb / ?pqda)) / (?pqdb + 1)"
    by (simp add: real_sqrt_mult [symmetric] real_sqrt_divide algebra_simps)
  also from ?pqdb = ?pqda * ?pqab and ?pqda > 0 and real_sqrt_pow2
  have " = (?pqda * sqrt ?pqab + sqrt ?pqab) / (?pqda * ?pqab + 1)"
    by (simp add: real_sqrt_mult power2_eq_square)
  finally
  have "?ac / ?bc = (?pqda * sqrt ?pqab + sqrt ?pqab) / (?pqda * ?pqab + 1)" .

  from ?pqda > 0 and ?pqab > 0
  have "?pqda * ?pqab + 1 > 0" by (simp add: add_pos_pos)
  with ?bc > 0
    and ?ac / ?bc = (?pqda * sqrt ?pqab + sqrt ?pqab) / (?pqda * ?pqab + 1)
  have "?ac * (?pqda * ?pqab + 1) = ?bc * (?pqda * sqrt ?pqab + sqrt ?pqab)"
    by (simp add: field_simps)
  hence "?pqda * (?ac * ?pqab - ?bc * sqrt ?pqab) = ?bc * sqrt ?pqab - ?ac"
    by (simp add: algebra_simps)

  from proj2_set_Col {p,q,a,b} and p  q and a  p and a  q
    and b  p
  have "cross_ratio_correct p q a b" by (unfold cross_ratio_correct_def) simp

  have "?ac * ?pqab - ?bc * sqrt ?pqab  0"
  proof
    assume "?ac * ?pqab - ?bc * sqrt ?pqab = 0"
    with ?pqda * (?ac * ?pqab - ?bc * sqrt ?pqab) = ?bc * sqrt ?pqab - ?ac
    have "?bc * sqrt ?pqab - ?ac = 0" by simp
    with ?ac * ?pqab - ?bc * sqrt ?pqab = 0 and ?ac > 0
    have "?pqab = 1" by simp
    with cross_ratio_correct p q a b
    have "a = b" by (rule cross_ratio_1_equal)
    with a  b show False ..
  qed
  with ?pqda * (?ac * ?pqab - ?bc * sqrt ?pqab) = ?bc * sqrt ?pqab - ?ac
  show "?pqda = (?bc * sqrt ?pqab - ?ac) / (?ac * ?pqab - ?bc * sqrt ?pqab)"
    by (simp add: field_simps)
qed

lemma perp_foot_cross_ratio_formula:
  assumes "a  hyp2" and "b  hyp2" and "c  hyp2" and "a  b"
  shows "cross_ratio (endpoint_in_S a b) (endpoint_in_S b a)
      (perp_foot c (proj2_line_through a b)) a
    = (cosh_dist b c * sqrt (exp_2dist a b) - cosh_dist a c)
      / (cosh_dist a c * exp_2dist a b - cosh_dist b c * sqrt (exp_2dist a b))"
  (is "cross_ratio ?p ?q ?d a
    = (?bc * sqrt ?pqab - ?ac) / (?ac * ?pqab - ?bc * sqrt ?pqab)")
proof -
  from a  b and a  hyp2 and b  hyp2
  have "are_endpoints_in_S ?p ?q a b"
    by (rule endpoints_in_S_are_endpoints_in_S)

  let ?l = "proj2_line_through a b"
  have "proj2_incident a ?l" and "proj2_incident b ?l"
    by (rule proj2_line_through_incident)+
  with a  b and a  hyp2 and b  hyp2
  have "proj2_incident ?p ?l" and "proj2_incident ?q ?l"
    by (simp_all add: endpoint_in_S_incident)

  let ?m = "drop_perp c ?l"
  have "M_perp ?l ?m" by (rule drop_perp_perp)

  have "proj2_incident ?d ?l" and "proj2_incident ?d ?m"
    by (rule perp_foot_incident)+

  have "proj2_incident c ?m" by (rule drop_perp_incident)
  with a  b and c  hyp2 and are_endpoints_in_S ?p ?q a b
    and proj2_incident ?p ?l and proj2_incident ?q ?l and M_perp ?l ?m
    and proj2_incident ?d ?l and proj2_incident ?d ?m
  have "cross_ratio ?p ?q ?d a
    = (?bc * sqrt (cross_ratio ?p ?q a b) - ?ac)
    / (?ac * (cross_ratio ?p ?q a b) - ?bc * sqrt (cross_ratio ?p ?q a b))"
    by (rule described_perp_foot_cross_ratio_formula)
  with a  b
  show "cross_ratio ?p ?q ?d a
    = (?bc * sqrt ?pqab - ?ac) / (?ac * ?pqab - ?bc * sqrt ?pqab)"
    by (unfold exp_2dist_def) simp
qed

subsection ‹The Klein--Beltrami model satisfies axiom 5›

lemma statement69:
  assumes "a b \<congruent>K a' b'" and "b c \<congruent>K b' c'" and "a c \<congruent>K a' c'"
  shows " J. is_K2_isometry J
   hyp2_cltn2 a J = a'  hyp2_cltn2 b J = b'  hyp2_cltn2 c J = c'"
proof cases
  assume "a = b"
  with a b \<congruent>K a' b' have "a' = b'" by (simp add: hyp2.A3_reversed)
  with a = b and b c \<congruent>K b' c'
  show " J. is_K2_isometry J
     hyp2_cltn2 a J = a'  hyp2_cltn2 b J = b'  hyp2_cltn2 c J = c'"
    by (unfold real_hyp2_C_def) simp
next
  assume "a  b"
  with a b \<congruent>K a' b'
  have "a'  b'" by (auto simp add: hyp2.A3')

  let ?pa = "Rep_hyp2 a"
    and ?pb = "Rep_hyp2 b"
    and ?pc = "Rep_hyp2 c"
    and ?pa' = "Rep_hyp2 a'"
    and ?pb' = "Rep_hyp2 b'"
    and ?pc' = "Rep_hyp2 c'"
  define pp pq l pp' pq' l'
    where "pp = endpoint_in_S ?pa ?pb"
      and "pq = endpoint_in_S ?pb ?pa"
      and "l = proj2_line_through ?pa ?pb"
      and "pp' = endpoint_in_S ?pa' ?pb'"
      and "pq' = endpoint_in_S ?pb' ?pa'"
      and "l' = proj2_line_through ?pa' ?pb'"
  define pd ps m pd' ps' m'
    where "pd = perp_foot ?pc l"
      and "ps = perp_up ?pc l"
      and "m = drop_perp ?pc l"
      and "pd' = perp_foot ?pc' l'"
      and "ps' = perp_up ?pc' l'"
      and "m' = drop_perp ?pc' l'"

  have "pp  S" and "pp'  S" and "pq  S" and "pq'  S"
    unfolding pp_def and pp'_def and pq_def and pq'_def
    by (simp_all add: Rep_hyp2 endpoint_in_S)

  from a  b and a'  b'
  have "?pa  ?pb" and "?pa'  ?pb'" by (unfold Rep_hyp2_inject)
  moreover
  have "proj2_incident ?pa l" and "proj2_incident ?pb l"
    and "proj2_incident ?pa' l'" and "proj2_incident ?pb' l'"
    by (unfold l_def l'_def) (rule proj2_line_through_incident)+
  ultimately have "proj2_incident pp l" and "proj2_incident pp' l'"
    and "proj2_incident pq l" and "proj2_incident pq' l'"
    unfolding pp_def and pp'_def and pq_def and pq'_def
    by (simp_all add: Rep_hyp2 endpoint_in_S_incident)

  from pp  S and pp'  S and proj2_incident pp l
    and proj2_incident pp' l' and proj2_incident ?pa l
    and proj2_incident ?pa' l'
  have "right_angle pp pd ps" and "right_angle pp' pd' ps'"
    unfolding pd_def and ps_def and pd'_def and ps'_def
    by (simp_all add: Rep_hyp2
      perp_foot_up_right_angle [of pp ?pc ?pa l]
      perp_foot_up_right_angle [of pp' ?pc' ?pa' l'])
  with right_angle_to_right_angle [of pp pd ps pp' pd' ps']
  obtain J where "is_K2_isometry J" and "apply_cltn2 pp J = pp'"
    and "apply_cltn2 pd J = pd'" and "apply_cltn2 ps J = ps'"
    by auto

  let ?paJ = "apply_cltn2 ?pa J"
    and ?pbJ = "apply_cltn2 ?pb J"
    and ?pcJ = "apply_cltn2 ?pc J"
    and ?pdJ = "apply_cltn2 pd J"
    and ?ppJ = "apply_cltn2 pp J"
    and ?pqJ = "apply_cltn2 pq J"
    and ?psJ = "apply_cltn2 ps J"
    and ?lJ = "apply_cltn2_line l J"
    and ?mJ = "apply_cltn2_line m J"

  have "proj2_incident pd l" and "proj2_incident pd' l'"
    and "proj2_incident pd m" and "proj2_incident pd' m'"
    by (unfold pd_def pd'_def m_def m'_def) (rule perp_foot_incident)+

  from proj2_incident pp l and proj2_incident pq l
    and proj2_incident pd l and proj2_incident ?pa l
    and proj2_incident ?pb l
  have "proj2_set_Col {pp,pq,pd,?pa}" and "proj2_set_Col {pp,pq,?pa,?pb}"
    by (unfold pd_def proj2_set_Col_def) (simp_all add: exI [of _ l])

  from ?pa  ?pb and ?pa'  ?pb'
  have "pp  pq" and "pp'  pq'"
    unfolding pp_def and pq_def and pp'_def and pq'_def
    by (simp_all add: Rep_hyp2 endpoint_in_S_swap)

  from proj2_incident ?pa l and proj2_incident ?pa' l'
  have "pd  hyp2" and "pd'  hyp2"
    unfolding pd_def and pd'_def
    by (simp_all add: Rep_hyp2 perp_foot_hyp2 [of ?pa l ?pc]
      perp_foot_hyp2 [of ?pa' l' ?pc'])

  from proj2_incident ?pa l and proj2_incident ?pa' l'
  have "ps  S" and "ps'  S"
    unfolding ps_def and ps'_def
    by (simp_all add: Rep_hyp2 perp_up_in_S [of ?pc ?pa l]
      perp_up_in_S [of ?pc' ?pa' l'])

  from pd  hyp2 and pp  S and ps  S
  have "pd  pp" and "?pa  pp" and "?pb  pp" and "pd  ps"
    by (simp_all add: Rep_hyp2 hyp2_S_not_equal)

  from is_K2_isometry J and pq  S
  have "?pqJ  S" by (unfold is_K2_isometry_def) simp

  from pd  pp and proj2_incident pd l and proj2_incident pp l
    and proj2_incident pd' l' and proj2_incident pp' l'
  have "?lJ = l'"
    unfolding ?pdJ = pd' [symmetric] and ?ppJ = pp' [symmetric]
    by (rule apply_cltn2_line_unique)
  from proj2_incident pq l and proj2_incident ?pa l
    and proj2_incident ?pb l
  have "proj2_incident ?pqJ l'" and "proj2_incident ?paJ l'"
    and "proj2_incident ?pbJ l'"
    by (unfold ?lJ = l' [symmetric]) simp_all

  from ?pa'  ?pb' and ?pqJ  S and proj2_incident ?pa' l'
    and proj2_incident ?pb' l' and proj2_incident ?pqJ l'
  have "?pqJ = pp'  ?pqJ = pq'"
    unfolding pp'_def and pq'_def
    by (simp add: Rep_hyp2 endpoints_in_S_incident_unique)
  moreover
  from pp  pq and apply_cltn2_injective
  have "pp'  ?pqJ" by (unfold ?ppJ = pp' [symmetric]) fast
  ultimately have "?pqJ = pq'" by simp

  from ?pa'  ?pb'
  have "cross_ratio pp' pq' pd' ?pa'
    = (cosh_dist ?pb' ?pc' * sqrt (exp_2dist ?pa' ?pb') - cosh_dist ?pa' ?pc')
      / (cosh_dist ?pa' ?pc' * exp_2dist ?pa' ?pb'
        - cosh_dist ?pb' ?pc' * sqrt (exp_2dist ?pa' ?pb'))"
    unfolding pp'_def and pq'_def and pd'_def and l'_def
    by (simp add: Rep_hyp2 perp_foot_cross_ratio_formula)
  also from assms
  have " = (cosh_dist ?pb ?pc * sqrt (exp_2dist ?pa ?pb) - cosh_dist ?pa ?pc)
    / (cosh_dist ?pa ?pc * exp_2dist ?pa ?pb
      - cosh_dist ?pb ?pc * sqrt (exp_2dist ?pa ?pb))"
    by (simp add: real_hyp2_C_exp_2dist real_hyp2_C_cosh_dist)
  also from ?pa  ?pb
  have " = cross_ratio pp pq pd ?pa"
    unfolding pp_def and pq_def and pd_def and l_def
    by (simp add: Rep_hyp2 perp_foot_cross_ratio_formula)
  also from proj2_set_Col {pp,pq,pd,?pa} and pp  pq and pd  pp
    and ?pa  pp
  have " = cross_ratio ?ppJ ?pqJ ?pdJ ?paJ" by (simp add: cross_ratio_cltn2)
  also from ?ppJ = pp' and ?pqJ = pq' and ?pdJ = pd'
  have " = cross_ratio pp' pq' pd' ?paJ" by simp
  finally
  have "cross_ratio pp' pq' pd' ?paJ = cross_ratio pp' pq' pd' ?pa'" by simp

  from is_K2_isometry J
  have "?paJ  hyp2" and "?pbJ  hyp2" and "?pcJ  hyp2"
    by (rule apply_cltn2_Rep_hyp2)+

  from proj2_incident pp' l' and proj2_incident pq' l'
    and proj2_incident pd' l' and proj2_incident ?paJ l'
    and proj2_incident ?pa' l' and proj2_incident ?pbJ l'
    and proj2_incident ?pb' l'
  have "proj2_set_Col {pp',pq',pd',?paJ}" and "proj2_set_Col {pp',pq',pd',?pa'}"
    and "proj2_set_Col {pp',pq',?pa',?pbJ}"
    and "proj2_set_Col {pp',pq',?pa',?pb'}"
    by (unfold proj2_set_Col_def) (simp_all add: exI [of _ l'])
  with pp'  pq' and pp'  S and pq'  S and pd'  hyp2
    and ?paJ  hyp2 and ?pbJ  hyp2
  have "are_endpoints_in_S pp' pq' pd' ?paJ"
    and "are_endpoints_in_S pp' pq' pd' ?pa'"
    and "are_endpoints_in_S pp' pq' ?pa' ?pbJ"
    and "are_endpoints_in_S pp' pq' ?pa' ?pb'"
    by (unfold are_endpoints_in_S_def) (simp_all add: Rep_hyp2)
  hence "cross_ratio_correct pp' pq' pd' ?paJ"
    and "cross_ratio_correct pp' pq' pd' ?pa'"
    and "cross_ratio_correct pp' pq' ?pa' ?pbJ"
    and "cross_ratio_correct pp' pq' ?pa' ?pb'"
    by (simp_all add: are_endpoints_in_S_cross_ratio_correct)

  from cross_ratio_correct pp' pq' pd' ?paJ
    and cross_ratio_correct pp' pq' pd' ?pa'
    and cross_ratio pp' pq' pd' ?paJ = cross_ratio pp' pq' pd' ?pa'
  have "?paJ = ?pa'" by (simp add: cross_ratio_unique)
  with ?ppJ = pp' and ?pqJ = pq'
  have "cross_ratio pp' pq' ?pa' ?pbJ = cross_ratio ?ppJ ?pqJ ?paJ ?pbJ" by simp
  also from proj2_set_Col {pp,pq,?pa,?pb} and pp  pq and ?pa  pp
    and ?pb  pp
  have " = cross_ratio pp pq ?pa ?pb" by (rule cross_ratio_cltn2)
  also from a  b and a b \<congruent>K a' b'
  have " = cross_ratio pp' pq' ?pa' ?pb'"
    unfolding pp_def pq_def pp'_def pq'_def
    by (rule real_hyp2_C_cross_ratio_endpoints_in_S)
  finally have "cross_ratio pp' pq' ?pa' ?pbJ = cross_ratio pp' pq' ?pa' ?pb'" .
  with cross_ratio_correct pp' pq' ?pa' ?pbJ
    and cross_ratio_correct pp' pq' ?pa' ?pb'
  have "?pbJ = ?pb'" by (rule cross_ratio_unique)

  let ?cc = "cart2_pt ?pc"
    and ?cd = "cart2_pt pd"
    and ?cs = "cart2_pt ps"
    and ?cc' = "cart2_pt ?pc'"
    and ?cd' = "cart2_pt pd'"
    and ?cs' = "cart2_pt ps'"
    and ?ccJ = "cart2_pt ?pcJ"
    and ?cdJ = "cart2_pt ?pdJ"
    and ?csJ = "cart2_pt ?psJ"

  from proj2_incident ?pa l and proj2_incident ?pa' l'
  have "B ?cd ?cc ?cs" and "B ?cd' ?cc' ?cs'"
    unfolding pd_def and ps_def and pd'_def and ps'_def
    by (simp_all add: Rep_hyp2 perp_up_at_end [of ?pc ?pa l]
      perp_up_at_end [of ?pc' ?pa' l'])

  from pd  hyp2 and ps  S and is_K2_isometry J
    and B ?cd ?cc ?cs
  have "B ?cdJ ?ccJ ?csJ" by (simp add: Rep_hyp2 statement_63)
  hence "B ?cd' ?ccJ ?cs'" by (unfold ?pdJ = pd' ?psJ = ps')

  from ?paJ = ?pa' have "cosh_dist ?pa' ?pcJ = cosh_dist ?paJ ?pcJ" by simp
  also from is_K2_isometry J
  have " = cosh_dist ?pa ?pc" by (simp add: Rep_hyp2 K2_isometry_cosh_dist)
  also from a c \<congruent>K a' c'
  have " = cosh_dist ?pa' ?pc'" by (rule real_hyp2_C_cosh_dist)
  finally have "cosh_dist ?pa' ?pcJ = cosh_dist ?pa' ?pc'" .

  have "M_perp l' m'" by (unfold m'_def) (rule drop_perp_perp)

  have "proj2_incident ?pc m" and "proj2_incident ?pc' m'"
    by (unfold m_def m'_def) (rule drop_perp_incident)+

  from proj2_incident ?pa l and proj2_incident ?pa' l'
  have "proj2_incident ps m" and "proj2_incident ps' m'"
    unfolding ps_def and m_def and ps'_def and m'_def
    by (simp_all add: Rep_hyp2 perp_up_incident [of ?pc ?pa l]
      perp_up_incident [of ?pc' ?pa' l'])
  with pd  ps and proj2_incident pd m and proj2_incident pd' m'
  have "?mJ = m'"
    unfolding ?pdJ = pd' [symmetric] and ?psJ = ps' [symmetric]
    by (simp add: apply_cltn2_line_unique)
  from proj2_incident ?pc m
  have "proj2_incident ?pcJ m'" by (unfold ?mJ = m' [symmetric]) simp
  with M_perp l' m' and Rep_hyp2 [of a'] and pd'  hyp2 and ?pcJ  hyp2
    and Rep_hyp2 [of c'] and proj2_incident ?pa' l'
    and proj2_incident pd' l' and proj2_incident pd' m'
    and proj2_incident ?pc' m'
  have "cosh_dist pd' ?pcJ = cosh_dist ?pa' ?pcJ / cosh_dist pd' ?pa'"
    and "cosh_dist pd' ?pc' = cosh_dist ?pa' ?pc' / cosh_dist pd' ?pa'"
    by (simp_all add: cosh_dist_perp_divide)
  with cosh_dist ?pa' ?pcJ = cosh_dist ?pa' ?pc'
  have "cosh_dist pd' ?pcJ = cosh_dist pd' ?pc'" by simp
  with pd'  hyp2 and ?pcJ  hyp2  and ?pc'  hyp2 and ps'  S
    and B ?cd' ?ccJ ?cs' and B ?cd' ?cc' ?cs'
  have "?pcJ = ?pc'" by (rule cosh_dist_unique)
  with ?paJ = ?pa' and ?pbJ = ?pb'
  have "hyp2_cltn2 a J = a'" and "hyp2_cltn2 b J = b'" and "hyp2_cltn2 c J = c'"
    by (unfold hyp2_cltn2_def) (simp_all add: Rep_hyp2_inverse)
  with is_K2_isometry J
  show " J. is_K2_isometry J
     hyp2_cltn2 a J = a'  hyp2_cltn2 b J = b'  hyp2_cltn2 c J = c'"
    by (simp add: exI [of _ J])
qed

theorem hyp2_axiom5:
  " a b c d a' b' c' d'.
  a  b  BK a b c  BK a' b' c'  a b \<congruent>K a' b'  b c \<congruent>K b' c'
     a d \<congruent>K a' d'  b d \<congruent>K b' d'
   c d \<congruent>K c' d'"
proof standard+
  fix a b c d a' b' c' d'
  assume "a  b  BK a b c  BK a' b' c'  a b \<congruent>K a' b'  b c \<congruent>K b' c'
     a d \<congruent>K a' d'  b d \<congruent>K b' d'"
  hence "a  b" and "BK a b c" and "BK a' b' c'" and "a b \<congruent>K a' b'"
    and "b c \<congruent>K b' c'" and "a d \<congruent>K a' d'" and "b d \<congruent>K b' d'"
    by simp_all

  from a b \<congruent>K a' b' and b d \<congruent>K b' d' and a d \<congruent>K a' d' and statement69 [of a b a' b' d d']
  obtain J where "is_K2_isometry J" and "hyp2_cltn2 a J = a'"
    and "hyp2_cltn2 b J = b'" and "hyp2_cltn2 d J = d'"
    by auto

  let ?aJ = "hyp2_cltn2 a J"
    and ?bJ = "hyp2_cltn2 b J"
    and ?cJ = "hyp2_cltn2 c J"
    and ?dJ = "hyp2_cltn2 d J"

  from a  b and a b \<congruent>K a' b'
  have "a'  b'" by (auto simp add: hyp2.A3')

  from is_K2_isometry J and BK a b c
  have "BK ?aJ ?bJ ?cJ" by (rule real_hyp2_B_hyp2_cltn2)
  hence "BK a' b' ?cJ" by (unfold ?aJ = a' ?bJ = b')

  from is_K2_isometry J
  have "b c \<congruent>K ?bJ ?cJ" by (rule real_hyp2_C_hyp2_cltn2)
  hence "b c \<congruent>K b' ?cJ" by (unfold ?bJ = b')
  from this and b c \<congruent>K b' c' have "b' ?cJ \<congruent>K b' c'" by (rule hyp2.A2')
  with a'  b' and BK a' b' ?cJ and BK a' b' c'
  have "?cJ = c'" by (rule hyp2_extend_segment_unique)
  from is_K2_isometry J
  show "c d \<congruent>K c' d'"
    unfolding ?cJ = c' [symmetric] and ?dJ = d' [symmetric]
    by (rule real_hyp2_C_hyp2_cltn2)
qed

interpretation hyp2: tarski_first5 real_hyp2_C real_hyp2_B
  using hyp2_axiom4 and hyp2_axiom5
  by unfold_locales

subsection ‹The Klein--Beltrami model satisfies axioms 6, 7, and 11›

theorem hyp2_axiom6: " a b. BK a b a  a = b"
proof standard+
  fix a b
  let ?ca = "cart2_pt (Rep_hyp2 a)"
    and ?cb = "cart2_pt (Rep_hyp2 b)"
  assume "BK a b a"
  hence "B ?ca ?cb ?ca" by (unfold real_hyp2_B_def hyp2_rep_def)
  hence "?ca = ?cb" by (rule real_euclid.A6')
  hence "Rep_hyp2 a = Rep_hyp2 b" by (simp add: Rep_hyp2 hyp2_S_cart2_inj)
  thus "a = b" by (unfold Rep_hyp2_inject)
qed

lemma between_inverse:
  assumes "B (hyp2_rep p) v (hyp2_rep q)"
  shows "hyp2_rep (hyp2_abs v) = v"
proof -
  let ?u = "hyp2_rep p"
  let ?w = "hyp2_rep q"
  have "norm ?u < 1" and "norm ?w < 1" by (rule norm_hyp2_rep_lt_1)+

  from B ?u v ?w
  obtain l where "l  0" and "l  1" and "v - ?u = l *R (?w - ?u)"
    by (unfold real_euclid_B_def) auto
  from v - ?u = l *R (?w - ?u)
  have "v = l *R ?w + (1 - l) *R ?u" by (simp add: algebra_simps)
  hence "norm v  norm (l *R ?w) + norm ((1 - l) *R ?u)"
    by (simp only: norm_triangle_ineq [of "l *R ?w" "(1 - l) *R ?u"])
  with l  0 and l  1
  have "norm v  l *R norm ?w + (1 - l) *R norm ?u" by simp

  have "norm v < 1"
  proof cases
    assume "l = 0"
    with v = l *R ?w + (1 - l) *R ?u
    have "v = ?u" by simp
    with norm ?u < 1 show "norm v < 1" by simp
  next
    assume "l  0"
    with norm ?w < 1 and l  0 have "l *R norm ?w < l" by simp

    with norm ?u < 1 and l  1
      and mult_mono [of "1 - l" "1 - l" "norm ?u" 1]
    have "(1 - l) *R norm ?u  1 - l" by simp
    with l *R norm ?w < l
    have "l *R norm ?w + (1 - l) *R norm ?u < 1" by simp
    with norm v  l *R norm ?w + (1 - l) *R norm ?u
    show "norm v < 1" by simp
  qed
  thus "hyp2_rep (hyp2_abs v) = v" by (rule hyp2_rep_abs)
qed

lemma between_switch:
  assumes "B (hyp2_rep p) v (hyp2_rep q)"
  shows "BK p (hyp2_abs v) q"
proof -
  from assms have "hyp2_rep (hyp2_abs v) = v" by (rule between_inverse)
  with assms show "BK p (hyp2_abs v) q" by (unfold real_hyp2_B_def) simp
qed

theorem hyp2_axiom7:
  " a b c p q. BK a p c  BK b q c  ( x. BK p x b  BK q x a)"
proof auto
  fix a b c p q
  let ?ca = "hyp2_rep a"
    and ?cb = "hyp2_rep b"
    and ?cc = "hyp2_rep c"
    and ?cp = "hyp2_rep p"
    and ?cq = "hyp2_rep q"
  assume "BK a p c" and "BK b q c"
  hence "B ?ca ?cp ?cc" and "B ?cb ?cq ?cc" by (unfold real_hyp2_B_def)
  with real_euclid.A7' [of ?ca ?cp ?cc ?cb ?cq]
  obtain cx where "B ?cp cx ?cb" and "B ?cq cx ?ca" by auto
  hence "BK p (hyp2_abs cx) b" and "BK q (hyp2_abs cx) a"
    by (simp_all add: between_switch)
  thus " x. BK p x b  BK q x a" by (simp add: exI [of _ "hyp2_abs cx"])
qed

theorem hyp2_axiom11:
  " X Y. ( a.  x y. x  X  y  Y  BK a x y)
   ( b.  x y. x  X  y  Y  BK x b y)"
proof (rule allI)+
  fix X Y :: "hyp2 set"
  show "( a.  x y. x  X  y  Y  BK a x y)
     ( b.  x y. x  X  y  Y  BK x b y)"
  proof cases
    assume "X = {}  Y = {}"
    thus "( a.  x y. x  X  y  Y  BK a x y)
       ( b.  x y. x  X  y  Y  BK x b y)" by auto
  next
    assume "¬ (X = {}  Y = {})"
    hence "X  {}" and "Y  {}" by simp_all
    then obtain w and z where "w  X" and "z  Y" by auto

    show "( a.  x y. x  X  y  Y  BK a x y)
       ( b.  x y. x  X  y  Y  BK x b y)"
    proof
      assume " a.  x y. x  X  y  Y  BK a x y"
      then obtain a where " x y. x  X  y  Y  BK a x y" ..

      let ?cX = "hyp2_rep ` X"
        and ?cY = "hyp2_rep ` Y"
        and ?ca = "hyp2_rep a"
        and ?cw = "hyp2_rep w"
        and ?cz = "hyp2_rep z"

      from  x y. x  X  y  Y  BK a x y
      have " cx cy. cx  ?cX  cy  ?cY  B ?ca cx cy"
        by (unfold real_hyp2_B_def) auto
      with real_euclid.A11' [of ?cX ?cY ?ca]
      obtain cb where " cx cy. cx  ?cX  cy  ?cY  B cx cb cy" by auto
      with w  X and z  Y have "B ?cw cb ?cz" by simp
      hence "hyp2_rep (hyp2_abs cb) = cb" (is "hyp2_rep ?b = cb")
        by (rule between_inverse)
      with  cx cy. cx  ?cX  cy  ?cY  B cx cb cy
      have " x y. x  X  y  Y  BK x ?b y"
        by (unfold real_hyp2_B_def) simp
      thus " b.  x y. x  X  y  Y  BK x b y" by (rule exI)
    qed
  qed
qed

interpretation tarski_absolute_space real_hyp2_C real_hyp2_B
  using hyp2_axiom6 and hyp2_axiom7 and hyp2_axiom11
  by unfold_locales

subsection ‹The Klein--Beltrami model satisfies the dimension-specific axioms›

lemma hyp2_rep_abs_examples:
  shows "hyp2_rep (hyp2_abs 0) = 0" (is "hyp2_rep ?a = ?ca")
  and "hyp2_rep (hyp2_abs (vector [1/2,0])) = vector [1/2,0]"
  (is "hyp2_rep ?b = ?cb")
  and "hyp2_rep (hyp2_abs (vector [0,1/2])) = vector [0,1/2]"
  (is "hyp2_rep ?c = ?cc")
  and "hyp2_rep (hyp2_abs (vector [1/4,1/4])) = vector [1/4,1/4]"
  (is "hyp2_rep ?d = ?cd")
  and "hyp2_rep (hyp2_abs (vector [1/2,1/2])) = vector [1/2,1/2]"
  (is "hyp2_rep ?t = ?ct")
proof -
  have "norm ?ca < 1" and "norm ?cb < 1" and "norm ?cc < 1" and "norm ?cd < 1"
    and "norm ?ct < 1"
    by (unfold norm_vec_def L2_set_def) (simp_all add: sum_2 power2_eq_square)
  thus "hyp2_rep ?a = ?ca" and "hyp2_rep ?b = ?cb" and "hyp2_rep ?c = ?cc"
    and "hyp2_rep ?d = ?cd" and "hyp2_rep ?t = ?ct"
    by (simp_all add: hyp2_rep_abs)
qed

theorem hyp2_axiom8: " a b c. ¬ BK a b c  ¬ BK b c a  ¬ BK c a b"
proof -
  let ?ca = "0 :: real^2"
    and ?cb = "vector [1/2,0] :: real^2"
    and ?cc = "vector [0,1/2] :: real^2"
  let ?a = "hyp2_abs ?ca"
    and ?b = "hyp2_abs ?cb"
    and ?c = "hyp2_abs ?cc"
  from hyp2_rep_abs_examples and non_Col_example
  have "¬ (hyp2.Col ?a ?b ?c)"
    by (unfold hyp2.Col_def real_euclid.Col_def real_hyp2_B_def) simp
  thus " a b c. ¬ BK a b c  ¬ BK b c a  ¬ BK c a b"
    unfolding hyp2.Col_def
    by simp (rule exI)+
qed

theorem hyp2_axiom9:
  " p q a b c. p  q  a p \<congruent>K a q  b p \<congruent>K b q  c p \<congruent>K c q
   BK a b c  BK b c a  BK c a b"
proof (rule allI)+
  fix p q a b c
  show "p  q  a p \<congruent>K a q  b p \<congruent>K b q  c p \<congruent>K c q
     BK a b c  BK b c a  BK c a b"
  proof
    assume "p  q  a p \<congruent>K a q  b p \<congruent>K b q  c p \<congruent>K c q"
    hence "p  q" and "a p \<congruent>K a q" and "b p \<congruent>K b q" and "c p \<congruent>K c q" by simp_all

    let ?pp = "Rep_hyp2 p"
      and ?pq = "Rep_hyp2 q"
      and ?pa = "Rep_hyp2 a"
      and ?pb = "Rep_hyp2 b"
      and ?pc = "Rep_hyp2 c"
    define l where "l = proj2_line_through ?pp ?pq"
    define m ps pt stpq
      where "m = drop_perp ?pa l"
        and "ps = endpoint_in_S ?pp ?pq"
        and "pt = endpoint_in_S ?pq ?pp"
        and "stpq = exp_2dist ?pp ?pq"

    from p  q have "?pp  ?pq" by (simp add: Rep_hyp2_inject)

    from Rep_hyp2
    have "stpq > 0" by (unfold stpq_def) (simp add: exp_2dist_positive)
    hence "sqrt stpq * sqrt stpq = stpq"
      by (simp add: real_sqrt_mult [symmetric])

    from Rep_hyp2 and ?pp  ?pq
    have "stpq  1" by (unfold stpq_def) (auto simp add: exp_2dist_1_equal)

    have "z_non_zero ?pa" and "z_non_zero ?pb" and "z_non_zero ?pc"
      by (simp_all add: Rep_hyp2 hyp2_S_z_non_zero)

    have " pd  {?pa,?pb,?pc}.
      cross_ratio ps pt (perp_foot pd l) ?pp = 1 / (sqrt stpq)"
    proof
      fix pd
      assume "pd  {?pa,?pb,?pc}"
      with Rep_hyp2 have "pd  hyp2" by auto

      define pe x
        where "pe = perp_foot pd l"
          and "x = cosh_dist ?pp pd"

      from pd  {?pa,?pb,?pc} and a p \<congruent>K a q and b p \<congruent>K b q
        and c p \<congruent>K c q
      have "cosh_dist pd ?pp = cosh_dist pd ?pq"
        by (auto simp add: real_hyp2_C_cosh_dist)
      with pd  hyp2 and Rep_hyp2
      have "x = cosh_dist ?pq pd" by (unfold x_def) (simp add: cosh_dist_swap)

      from Rep_hyp2 [of p] and pd  hyp2 and cosh_dist_positive [of ?pp pd]
      have "x  0" by (unfold x_def) simp

      from Rep_hyp2 and pd  hyp2 and ?pp  ?pq
      have "cross_ratio ps pt pe ?pp
        = (cosh_dist ?pq pd * sqrt stpq - cosh_dist ?pp pd)
        / (cosh_dist ?pp pd *  stpq - cosh_dist ?pq pd * sqrt stpq)"
        unfolding ps_def and pt_def and pe_def and l_def and stpq_def
        by (simp add: perp_foot_cross_ratio_formula)
      also from x_def and x = cosh_dist ?pq pd
      have " = (x * sqrt stpq - x) / (x * stpq - x * sqrt stpq)" by simp
      also from sqrt stpq * sqrt stpq = stpq
      have " = (x * sqrt stpq - x) / ((x * sqrt stpq - x) * sqrt stpq)"
        by (simp add: algebra_simps)
      also from x  0 and stpq  1 have " = 1 / sqrt stpq" by simp
      finally show "cross_ratio ps pt pe ?pp = 1 / sqrt stpq" .
    qed
    hence "cross_ratio ps pt (perp_foot ?pa l) ?pp = 1 / sqrt stpq" by simp

    have " pd  {?pa,?pb,?pc}. proj2_incident pd m"
    proof
      fix pd
      assume "pd  {?pa,?pb,?pc}"
      with Rep_hyp2 have "pd  hyp2" by auto
      with Rep_hyp2 and ?pp  ?pq and proj2_line_through_incident
      have "cross_ratio_correct ps pt ?pp (perp_foot pd l)"
        and "cross_ratio_correct ps pt ?pp (perp_foot ?pa l)"
        unfolding ps_def and pt_def and l_def
        by (simp_all add: endpoints_in_S_perp_foot_cross_ratio_correct)

      from pd  {?pa,?pb,?pc}
        and  pd  {?pa,?pb,?pc}.
        cross_ratio ps pt (perp_foot pd l) ?pp = 1 / (sqrt stpq)
      have "cross_ratio ps pt (perp_foot pd l) ?pp = 1 / sqrt stpq" by auto
      with cross_ratio ps pt (perp_foot ?pa l) ?pp = 1 / sqrt stpq
      have "cross_ratio ps pt (perp_foot pd l) ?pp
        = cross_ratio ps pt (perp_foot ?pa l) ?pp"
        by simp
      hence "cross_ratio ps pt ?pp (perp_foot pd l)
        = cross_ratio ps pt ?pp (perp_foot ?pa l)"
        by (simp add: cross_ratio_swap_34 [of ps pt _ ?pp])
      with cross_ratio_correct ps pt ?pp (perp_foot pd l)
        and cross_ratio_correct ps pt ?pp (perp_foot ?pa l)
      have "perp_foot pd l = perp_foot ?pa l" by (rule cross_ratio_unique)
      with Rep_hyp2 [of p] and pd  hyp2
        and proj2_line_through_incident [of ?pp ?pq]
        and perp_foot_eq_implies_drop_perp_eq [of ?pp pd l ?pa]
      have "drop_perp pd l = m" by (unfold m_def l_def) simp
      with drop_perp_incident [of pd l] show "proj2_incident pd m" by simp
    qed
    hence "proj2_set_Col {?pa,?pb,?pc}"
      by (unfold proj2_set_Col_def) (simp add: exI [of _ m])
    hence "proj2_Col ?pa ?pb ?pc" by (simp add: proj2_Col_iff_set_Col)
    with z_non_zero ?pa and z_non_zero ?pb and z_non_zero ?pc
    have "real_euclid.Col (hyp2_rep a) (hyp2_rep b) (hyp2_rep c)"
      by (unfold hyp2_rep_def) (simp add: proj2_Col_iff_euclid_cart2)
    thus "BK a b c  BK b c a  BK c a b"
      by (unfold real_hyp2_B_def real_euclid.Col_def)
  qed
qed

interpretation hyp2: tarski_absolute real_hyp2_C real_hyp2_B
  using hyp2_axiom8 and hyp2_axiom9
  by unfold_locales

subsection ‹The Klein--Beltrami model violates the Euclidean axiom›

theorem hyp2_axiom10_false:
  shows "¬ ( a b c d t. BK a d t  BK b d c  a  d
   ( x y. BK a b x  BK a c y  BK x t y))"
proof
  assume " a b c d t. BK a d t  BK b d c  a  d
     ( x y. BK a b x  BK a c y  BK x t y)"

  let ?ca = "0 :: real^2"
    and ?cb = "vector [1/2,0] :: real^2"
    and ?cc = "vector [0,1/2] :: real^2"
    and ?cd = "vector [1/4,1/4] :: real^2"
    and ?ct = "vector [1/2,1/2] :: real^2"
  let ?a = "hyp2_abs ?ca"
    and ?b = "hyp2_abs ?cb"
    and ?c = "hyp2_abs ?cc"
    and ?d = "hyp2_abs ?cd"
    and ?t = "hyp2_abs ?ct"

  have "?cd = (1/2) *R ?ct" and "?cd - ?cb = (1/2) *R (?cc - ?cb)"
    by (unfold vector_def) (simp_all add: vec_eq_iff)
  hence "B ?ca ?cd ?ct" and "B ?cb ?cd ?cc"
    by (unfold real_euclid_B_def) (simp_all add: exI [of _ "1/2"])
  hence "BK ?a ?d ?t" and "BK ?b ?d ?c"
    by (unfold real_hyp2_B_def) (simp_all add: hyp2_rep_abs_examples)

  have "?a  ?d"
  proof
    assume "?a = ?d"
    hence "hyp2_rep ?a = hyp2_rep ?d" by simp
    hence "?ca = ?cd" by (simp add: hyp2_rep_abs_examples)
    thus False by (simp add: vec_eq_iff forall_2)
  qed
  with BK ?a ?d ?t and BK ?b ?d ?c
    and  a b c d t. BK a d t  BK b d c  a  d
     ( x y. BK a b x  BK a c y  BK x t y)
  obtain x and y where "BK ?a ?b x" and "BK ?a ?c y" and "BK x ?t y"
    by blast

  let ?cx = "hyp2_rep x"
    and ?cy = "hyp2_rep y"
  from BK ?a ?b x and BK ?a ?c y and BK x ?t y
  have "B ?ca ?cb ?cx" and "B ?ca ?cc ?cy" and "B ?cx ?ct ?cy"
    by (unfold real_hyp2_B_def) (simp_all add: hyp2_rep_abs_examples)

  from B ?ca ?cb ?cx and B ?ca ?cc ?cy and B ?cx ?ct ?cy
  obtain j and k and l where "?cb - ?ca = j *R (?cx - ?ca)"
    and "?cc - ?ca = k *R (?cy - ?ca)"
    and "l  0" and "l  1" and "?ct - ?cx = l *R (?cy - ?cx)"
    by (unfold real_euclid_B_def) fast

  from ?cb - ?ca = j *R (?cx - ?ca) and ?cc - ?ca = k *R (?cy - ?ca)
  have "j  0" and "k  0" by (auto simp add: vec_eq_iff forall_2)
  with ?cb - ?ca = j *R (?cx - ?ca) and ?cc - ?ca = k *R (?cy - ?ca)
  have "?cx = (1/j) *R ?cb" and "?cy = (1/k) *R ?cc" by simp_all
  hence "?cx$2 = 0" and "?cy$1 = 0" by simp_all

  from ?ct - ?cx = l *R (?cy - ?cx)
  have "?ct = (1 - l) *R ?cx + l *R ?cy" by (simp add: algebra_simps)
  with ?cx$2 = 0 and ?cy$1 = 0
  have "?ct$1 = (1 - l) * (?cx$1)" and "?ct$2 = l * (?cy$2)" by simp_all
  hence "l * (?cy$2) = 1/2" and "(1 - l) * (?cx$1) = 1/2" by simp_all

  have "?cx$1  ¦?cx$1¦" by simp
  also have "  norm ?cx" by (rule component_le_norm_cart)
  also have " < 1" by (rule norm_hyp2_rep_lt_1)
  finally have "?cx$1 < 1" .
  with l  1 and mult_less_cancel_left [of "1 - l" "?cx$1" 1]
  have "(1 - l) * ?cx$1  1 - l" by auto
  with (1 - l) * (?cx$1) = 1/2 have "l  1/2" by simp

  have "?cy$2  ¦?cy$2¦" by simp
  also have "  norm ?cy" by (rule component_le_norm_cart)
  also have " < 1" by (rule norm_hyp2_rep_lt_1)
  finally have "?cy$2 < 1" .
  with l  0 and mult_less_cancel_left [of l "?cy$2" 1]
  have "l * ?cy$2  l" by auto
  with l * (?cy$2) = 1/2 have "l  1/2" by simp
  with l  1/2 have "l = 1/2" by simp
  with l * (?cy$2) = 1/2 have "?cy$2 = 1" by simp
  with ?cy$2 < 1 show False by simp
qed

theorem hyp2_not_tarski: "¬ (tarski real_hyp2_C real_hyp2_B)"
  using hyp2_axiom10_false
  by (unfold tarski_def tarski_space_def tarski_space_axioms_def) simp

text ‹Therefore axiom 10 is independent.›

end