Theory Tarski_Neutral_Archimedes

(* IsaGeoCoq - Tarski_Neutral_Archimedes.thy

Port part of GeoCoq 3.4.0 (https://geocoq.github.io/GeoCoq/) in Isabelle/Hol 

Copyright (C) 2021-2025  Roland Coghetto roland.coghetto (at) cafr-msa2p.be

License: LGPL

This library is free software; you can redistribute it and/or
modify it under the terms of the GNU Lesser General Public
License as published by the Free Software Foundation; either
version 2.1 of the License, or (at your option) any later version.

This library is distributed in the hope that it will be useful,
but WITHOUT ANY WARRANTY; without even the implied warranty of
MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE.  See the GNU
Lesser General Public License for more details.

You should have received a copy of the GNU Lesser General Public
License along with this library; if not, write to the Free Software
Foundation, Inc., 51 Franklin Street, Fifth Floor, Boston, MA  02110-1301  USA
*)

theory Tarski_Neutral_Archimedes

imports 
  Tarski_Neutral

begin

section "Continuity Axioms"

context Tarski_neutral_dimensionless

begin

subsection "Definitions"

definition greenberg_s_axiom ::
  "bool"
  ("GreenBergsAxiom")
  where
    "greenberg_s_axiom   P Q R A B C. 
 ¬ Col A B C  Acute A B C  Q  R  Per P Q R  ( S. P S Q LtA A B C  Q Out S R)"

definition aristotle_s_axiom ::
  "bool"
  ("AristotleAxiom") where
  "aristotle_s_axiom   P Q A B C.
¬ Col A B C  Acute A B C 
( X Y. B Out A X  B Out C Y  Per B X Y  P Q Lt X Y)"

definition Axiom1:: "bool" where "Axiom1   A B C D. 
( I. Col I A B  Col I C D)  ¬  ( I. Col I A B  Col I C D)"

definition PreGrad :: "'p  'p  'p  'p  bool" where
  "PreGrad A B C D  (A  B  Bet A B C  Bet A C D  Cong A B C D)"

fun Sym :: "'p  'p  'p  'p" where
  "Sym A B C = (if (A  B  Bet A B C) then
      (SOME x::'p. PreGrad A B C x)
         else
       A)"

fun Gradn :: "['p,'p]  nat 'p"where
  "Gradn A B n = (if (A = B) then 
                    A
                      else 
                    (if (n = 0) then 
                       A 
                         else 
                     (if (n = 1) then 
                        B 
                          else 
                        (Sym A B (Gradn A B (n-1))))))"

definition Grad :: "['p,'p,'p]  bool" where
  "Grad A B C   n. (n  0)  (C = Gradn A B n)"

inductive GradI :: "['p,'p,'p]  bool" for A B
  where
    gradi_init : "GradI A B B" 
  | gradi_stab : "GradI A B C'" if 
    "GradI A B C" 
    and "Bet A C C'" 
    and "Cong A B C C'" 

definition Reach :: "['p,'p,'p,'p]  bool" where
  "Reach A B C D   B'. Grad A B B'  C D Le A B'"

definition archimedes_axiom ::
  "bool"
  ("ArchimedesAxiom") where
  "archimedes_axiom   A B C D::'p. 
A  B  Reach A B C D"

inductive GradA :: "['p,'p,'p,'p,'p,'p]  bool" for A B C
  where
    grada_init : "GradA A B C D E F" if 
    "A B C CongA D E F" 
  | grada_stab : "GradA A B C G H I" if 
    "GradA A B C D E F" 
    and "SAMS D E F A B C" 
    and "D E F A B C SumA G H I" 

inductive GradAExp :: "['p,'p,'p,'p,'p,'p]  bool" for A B C 
  where
    gradaexp_init : "GradAExp A B C D E F" if 
    "A B C CongA D E F" 
  | gradaexp_stab : "GradAExp A B C G H I" if 
    "GradAExp A B C D E F" 
    and "SAMS D E F D E F" 
    and "D E F D E F SumA G H I" 

definition Grad2 :: "['p,'p,'p,'p,'p,'p]  bool" where
  "Grad2 A B C D E F   n. (n  0)  (C = Gradn A B n)  (F = Gradn D E n)"

fun SymR :: "'p  'p  'p" where
  "SymR A B = (SOME x::'p. B Midpoint A x)"

fun GradExpn :: "'p  'p  nat  'p" where
  "(GradExpn A B n) = (if (A = B) then 
                         A 
                           else
                         (if (n = 0) then 
                            A
                              else
                            (if (n = 1) then 
                               B
                                 else 
                               (SymR A (GradExpn A B (n-1))))))"

definition GradExp :: "'p  'p  'p  bool" where
  "GradExp A B C    n. (n  0)  C = GradExpn A B n"

definition GradExp2 :: "['p,'p,'p,'p,'p,'p]  bool" where
  "GradExp2 A B C D E F   n. (n  0)  (C = GradExpn A B n)  (F = GradExpn D E n)"

fun MidR :: "'p  'p  'p" where
  "MidR A B = (SOME x. x Midpoint A B)"

fun GradExpInvn :: "'p  'p  nat  'p" where
  "(GradExpInvn A B n) = (if (A = B) then 
                            A 
                              else
                            (if (n = 0) then 
                               B
                                 else
                               (if (n = 1) then 
                                  (MidR A B)
                                    else 
                                  (MidR A (GradExpInvn A B (n-1))))))"

definition GradExpInv :: "'p  'p  'p  bool" where
  "GradExpInv A B C    n. B = GradExpInvn A C n"

subsection "Propositions"

lemma PreGrad_lem1:
  assumes "A  B" and 
    "Bet A B C"
  shows " x. PreGrad A B C x"
  by (meson PreGrad_def assms(1) assms(2) not_cong_3412 segment_construction) 

lemma PreGrad_uniq:
  assumes "PreGrad A B C x" and
    "PreGrad A B C y"
  shows "x = y"
  by (metis (no_types, lifting) PreGrad_def assms(1) assms(2) 
      bet_neq12__neq between_cong_3 cong_inner_transitivity)

lemma Diff_Mid__PreGrad:
  assumes "A  B" and
    "B Midpoint A C"
  shows "PreGrad A B B C"
  by (simp add: PreGrad_def assms(1) assms(2) between_trivial midpoint_bet midpoint_cong)

lemma Diff_Mid_Mid_PreGrad:
  assumes "A  B" and
    "B Midpoint A C" and
    "C Midpoint B D" 
  shows "PreGrad A B C D"
proof -
  have "Bet A B C" 
    using Midpoint_def assms(2) by presburger
  moreover have "Bet A C D" 
    using Midpoint_def assms(3) calculation is_midpoint_id outer_transitivity_between2 by blast
  moreover have "Cong A B C D" 
    using assms(2) assms(3) cong_transitivity midpoint_cong by blast
  ultimately show ?thesis 
    by (simp add: assms(1) PreGrad_def)
qed

lemma Sym_Diff__Diff:
  assumes "Sym A B C = D" and 
    "A  D"
  shows "A  B"
  using assms(1) assms(2) by force 

lemma Sym_Refl:
  "Sym A A A = A"
  by simp

lemma Diff_Mid__Sym:
  assumes "A  B" and
    "B Midpoint A C"
  shows "Sym A B B = C" 
  using someI_ex by (metis Diff_Mid__PreGrad Sym.elims PreGrad_uniq 
      assms(1) assms(2) between_trivial) 

lemma Mid_Mid__Sym:
  assumes "A  B" and
    "B Midpoint A C" and
    "C Midpoint B D"
  shows "Sym A B C = D" 
proof -
  have "PreGrad A B C D"
    by (simp add: Diff_Mid_Mid_PreGrad assms(1) assms(2) assms(3))
  thus ?thesis 
    using someI_ex assms(1) 
    by (metis PreGrad_uniq Sym.elims assms(2) midpoint_bet) 
qed

lemma Sym_Bet__Bet_Bet:
  assumes "Sym A B C = D" and 
    "A  B" and 
    "Bet A B C" 
  shows "Bet A B D  Bet A C D"
proof -
  have "(SOME x::'p. PreGrad A B C x) = D"
    using assms(1) assms(2) assms(3) by auto
  hence "PreGrad A B C D"
    by (metis PreGrad_lem1 assms(2) assms(3) someI2) 
  thus ?thesis
    by (meson PreGrad_def between_exchange4) 
qed

lemma Sym_Bet__Cong:
  assumes "Sym A B C = D" and 
    "A  B" and 
    "Bet A B C" 
  shows "Cong A B C D"
proof -
  have "(SOME x::'p. PreGrad A B C x) = D"
    using assms(1) assms(2) assms(3) by auto
  hence "PreGrad A B C D"
    by (metis PreGrad_lem1 assms(2) assms(3) someI2) 
  thus ?thesis
    by (meson PreGrad_def between_exchange4) 
qed

lemma LemSym_aux:
  assumes "A  B" and 
    "Bet A B C" and
    "Bet A C D" and
    "Cong A B C D"
  shows "Sym A B C = D"
proof -
  have "PreGrad A B C D"
    using PreGrad_def assms(1) assms(2) assms(3) assms(4) by blast
  thus ?thesis
    by (metis PreGrad_def PreGrad_uniq Sym_Bet__Bet_Bet Sym_Bet__Cong)
qed

lemma Lem_Gradn_id_n:
  "Gradn A A n = A" 
  by simp

lemma Lem_Gradn_0:
  "Gradn A B 0 = A" 
  by simp

lemma Lem_Gradn_1:
  "Gradn A B 1 = B" 
  by simp

lemma Diff__Gradn_Sym:
  assumes "A  B" and
    "n > 1"
  shows "Gradn A B n = Sym A B (Gradn A B (n-1))" 
proof -
  have "¬ (n = 0  n = 1)"
    by auto
  thus ?thesis 
    using assms(1) assms(2) by simp
qed

lemma Diff__Bet_Gradn_Suc:
  assumes "A  B"
  shows "Bet A B (Gradn A B (Suc n))"
proof (induction n)
  case 0
  hence "Gradn A B (Suc 0) = B" 
    using assms(1) by simp
  thus ?case
    by (simp add: between_trivial) 
next
  case (Suc n)
  {
    assume 1: "Bet A B (Gradn A B (Suc n))"
    have "Gradn A B (Suc (Suc n)) = Sym A B (Gradn A B (Suc n))"
      by simp
    hence "Bet A B (Gradn A B (Suc (Suc n)))  
    Bet A (Gradn A B (Suc n)) (Gradn A B (Suc (Suc n)))"
      using 1 Sym_Bet__Bet_Bet assms by presburger 
    hence "Bet A B (Gradn A B (Suc (Suc n)))"
      by simp
  }
  thus ?case
    using Suc.IH by blast 
qed

lemma Diff_Le_Gradn_Suc:
  assumes "A  B"
  shows "A B Le A (Gradn A B (Suc n))"
  by (meson Diff__Bet_Gradn_Suc assms bet__le1213)

lemma Diff__Bet_Gradn:
  assumes "A  B" and
    "n  0"
  shows "Bet A B (Gradn A B n)" 
  using assms(1) assms(2) Diff__Bet_Gradn_Suc not0_implies_Suc by blast

lemma Diff_Le_Gradn_n:
  assumes "A  B" and
    "n  0"
  shows "A B Le A (Gradn A B n)"
  by (meson Diff__Bet_Gradn assms(1) assms(2) l5_12_a)

lemma Diff_Bet_Gradn_Suc_Gradn_Suc2:
  assumes "A  B"
  shows "Bet A (Gradn A B (Suc n)) (Gradn A B (Suc (Suc n)))"
proof (induction n)
  case 0
  hence 1: "Gradn A B (Suc 0) = B" 
    using assms(1) by simp
  from assms(1) 
  have "(Gradn A B (Suc (Suc 0))) = (Sym A B (Gradn A B (Suc 0)))" 
    by simp
  thus ?case
    by (metis "1" Diff__Bet_Gradn_Suc assms) 
next
  case (Suc n)
  {
    assume 1: "Bet A (Gradn A B (Suc n)) (Gradn A B (Suc (Suc n)))"
    have "Gradn A B (Suc (Suc n)) = Sym A B (Gradn A B (Suc n))"
      by simp
    hence "Bet A B (Gradn A B (Suc (Suc n)))"
      using Diff__Bet_Gradn_Suc assms by blast 
    have "Gradn A B (Suc(Suc (Suc n))) = Sym A B (Gradn A B (Suc(Suc n)))"
      by simp
    hence "PreGrad A B (Gradn A B (Suc(Suc n))) (Gradn A B (Suc(Suc (Suc n))))"
      using PreGrad_def Sym_Bet__Bet_Bet Sym_Bet__Cong 
        Bet A B (Gradn A B (Suc (Suc n))) assms by presburger
    hence "Bet A B (Gradn A B (Suc(Suc (Suc n))))  
    Bet A (Gradn A B (Suc(Suc n))) (Gradn A B (Suc(Suc (Suc n))))"
      by (metis Diff__Bet_Gradn_Suc Sym_Bet__Bet_Bet 
          Gradn A B (Suc (Suc (Suc n))) = Sym A B (Gradn A B (Suc (Suc n))) assms)     
    hence "Bet A (Gradn A B (Suc(Suc n))) (Gradn A B (Suc(Suc (Suc n))))"
      by blast
  }
  thus ?case
    using Suc.IH by blast
qed

lemma Diff__Bet_Gradn_Gradn_SucA:
  assumes "A  B"
  shows "A (Gradn A B (Suc n)) Le A (Gradn A B (Suc (Suc n)))"
  by (meson Diff_Bet_Gradn_Suc_Gradn_Suc2 assms bet__le1213)

lemma Diff__Bet_Gradn_Gradn_Suc:
  assumes "A  B"
  shows "Bet A (Gradn A B n) (Gradn A B (Suc n))"
proof (induction n)
  case 0
  hence "Gradn A B 0 = A" by simp
  thus ?case
    using between_trivial2 by presburger  
next
  case (Suc n)
  {
    assume 1: "Bet A (Gradn A B n) (Gradn A B (Suc n))"
    have "Gradn A B (Suc (Suc n)) = Sym A B (Gradn A B (Suc n))"
      by simp
    hence "Bet A B (Gradn A B (Suc (Suc n)))  
    Bet A (Gradn A B (Suc n)) (Gradn A B (Suc (Suc n)))"
      using 1 Sym_Bet__Bet_Bet assms Diff__Bet_Gradn_Suc by presburger 
    hence "Bet A (Gradn A B (Suc n)) (Gradn A B (Suc (Suc n)))"
      by blast
  }
  thus ?case
    using Suc.IH by simp
qed

lemma Bet_Gradn_Gradn_Suc: 
  shows "Bet A (Gradn A B n) (Gradn A B (Suc n))"
  by (metis Lem_Gradn_id_n Diff__Bet_Gradn_Gradn_Suc not_bet_distincts)

lemma Gradn_Le_Gradn_Suc: 
  shows "A (Gradn A B n) Le A (Gradn A B (Suc n))"
  using Bet_Gradn_Gradn_Suc bet__le1213 by blast

lemma Bet_Gradn_Suc_Gradn_Suc2:
  shows "Bet B (Gradn A B (Suc n)) (Gradn A B (Suc(Suc n)))"
  by (metis Bet_Gradn_Gradn_Suc Diff__Bet_Gradn_Suc between_exchange3) 

lemma Gradn_Suc_Le_Gradn_Suc2:
  shows "B (Gradn A B (Suc n)) Le B (Gradn A B (Suc(Suc n)))"
  using Bet_Gradn_Suc_Gradn_Suc2 bet__le1213 by blast

lemma Diff_Le__Bet_Gradn_Plus:
  assumes "A  B" and
    "n  m"
  shows "Bet A (Gradn A B n) (Gradn A B (k + n))"
proof (induction k)
  case 0
  thus ?case
    using between_trivial by auto
next
  case (Suc k)
  {
    assume "Bet A (Gradn A B n) (Gradn A B (k + n))" 
    have "Bet A (Gradn A B (k + n)) (Gradn A B (Suc (k + n)))"
      using Diff__Bet_Gradn_Gradn_Suc assms(1) by presburger 
    hence "Bet A (Gradn A B n) (Gradn A B ((Suc k) + n))"
      by (metis Bet A (Gradn A B n) (Gradn A B (k + n)) 
          add_Suc between_exchange4) 
  }
  thus ?case
    using Suc.IH by blast 
qed

lemma  Diff_Le_Gradn_Plus:
  assumes "A  B" and
    "n  m"
  shows "A (Gradn A B n) Le A (Gradn A B (k + n))"
  by (meson Diff_Le__Bet_Gradn_Plus assms(1) assms(2) l5_12_a)

lemma Diff_Le_Bet__Gradn_Gradn:
  assumes "A  B" and
    "n  m"
  shows "Bet A (Gradn A B n) (Gradn A B m)"
proof (cases "n = 0")
  case True
  thus ?thesis
    using Lem_Gradn_0 between_trivial2 by presburger 
next
  case False
  hence 1: "n  0" 
    by auto
  show "Bet A (Gradn A B n) (Gradn A B m)"
  proof (cases "n = m")
    case True
    thus ?thesis 
      using between_trivial by presburger
  next
    case False
    hence "n < m"
      using assms(2) le_neq_implies_less by blast
    then obtain k where "m = k + n"
      using add.commute assms(2) le_Suc_ex by blast 
    have "Bet A (Gradn A B n) (Gradn A B (k + n))"
      using Diff_Le__Bet_Gradn_Plus assms(1) by blast 
    thus ?thesis
      using m = k + n by blast 
  qed
qed

lemma Diff_Le_Gradn:
  assumes "A  B" and
    "n  m"
  shows "A (Gradn A B n) Le A (Gradn A B m)"
  by (metis Diff_Le_Bet__Gradn_Gradn bet__le1213 assms(1) assms(2))

lemma Diff__Cong_Gradn_Suc_Gradn_Suc2:
  assumes "A  B"
  shows "Cong A B (Gradn A B (Suc n)) (Gradn A B  (Suc (Suc n)))"
proof (induction n)
  case 0
  hence 1: "Gradn A B (Suc 0) = B" 
    using assms(1) by simp
  from assms(1) 
  have "(Gradn A B (Suc (Suc 0))) = (Sym A B (Gradn A B (Suc 0)))" by simp
  hence "(Gradn A B (Suc (Suc 0))) = (Sym A B B)"
    using "1" by presburger 
  obtain C where "B Midpoint A C"
    using symmetric_point_construction by blast
  hence "C = Sym A B B"
    using Diff_Mid__Sym assms by blast
  hence "(Gradn A B (Suc (Suc 0))) = C"
    using "1" Gradn A B (Suc (Suc 0)) = Sym A B (Gradn A B (Suc 0)) by presburger 
  have "Cong A B (Gradn A B (Suc 0)) (Gradn A B  (Suc (Suc 0)))"
    using "1" B Midpoint A C Gradn A B (Suc (Suc 0)) = C midpoint_cong 
    by presburger 
  thus ?case
    by blast 
next
  case (Suc n)
  {
    assume "Cong A B (Gradn A B (Suc n)) (Gradn A B  (Suc (Suc n)))"
    have 1: "Gradn A B (Suc(Suc (Suc n))) = Sym A B (Gradn A B (Suc(Suc n)))"
      by simp
    have "Bet A B (Gradn A B (Suc (Suc n)))"
      using Diff__Bet_Gradn_Suc assms by blast
    hence "PreGrad A B (Gradn A B (Suc(Suc n))) (Gradn A B (Suc(Suc (Suc n))))"
      using 1 assms 
      by (metis PreGrad_def Sym_Bet__Cong Diff_Bet_Gradn_Suc_Gradn_Suc2) 
    hence "Cong A B (Gradn A B (Suc (Suc n))) (Gradn A B  (Suc(Suc (Suc n))))"
      using "1" Sym_Bet__Cong Bet A B (Gradn A B (Suc (Suc n))) assms 
      by presburger
  }
  thus ?case
    using Suc.IH by blast 
qed

lemma Cong_Gradn_Suc_Gradn_Suc2:
  shows "Cong A B (Gradn A B (Suc n)) (Gradn A B (Suc (Suc n)))"
  using Diff__Cong_Gradn_Suc_Gradn_Suc2 cong_reflexivity by auto 

lemma Cong_Gradn_Gradn_Suc:
  shows "Cong a b (Gradn a b n) (Gradn a b (Suc n))"
proof (cases "a = b")
  case True
  thus ?thesis
    by (simp add: cong_trivial_identity)
next
  case False
  have 1: "Gradn a b 0 = a"
    by auto
  have 2: "(Gradn a b (Suc 0)) = b"
    by auto
  {
    assume 3: "n = 0"
    hence "Cong a b a b  (Suc n = 1)"
      by (simp add: cong_reflexivity) 
    hence "Cong a b (Gradn a b n) (Gradn a b (Suc n))" 
      using 1 2 3 by presburger
  }
  moreover
  {
    assume "n  0"
    hence "Cong a b (Gradn a b n) (Gradn a b (Suc n))"
      by (metis False Diff__Cong_Gradn_Suc_Gradn_Suc2 old.nat.exhaust) 
  }
  ultimately
  show ?thesis
    by blast
qed

lemma Diff_Bet_Bet_Cong_Gradn_Suc: 
  assumes "A  B" and 
    "Bet A B C" and
    "Bet A (Gradn A B n) C" and
    "Cong A B (Gradn A B n) C"
  shows "C = (Gradn A B (Suc n))"
proof (cases "n = 0")
  case True
  thus ?thesis
    by (metis Lem_Gradn_0 Diff__Bet_Gradn_Suc Cong_Gradn_Gradn_Suc 
        assms(1) assms(2) assms(4) between_cong)  
next
  case False
  hence "Bet A B (Gradn A B n)"
    using Diff__Bet_Gradn assms(1) by blast 
  thus ?thesis
    by (metis LemSym_aux Diff__Bet_Gradn_Gradn_Suc Cong_Gradn_Gradn_Suc 
        assms(1) assms(3) assms(4)) 
qed

lemma grad_rec_0_1:
  shows "Cong a b (Gradn a b 0) (Gradn a b 1)"
  by (simp add: cong_reflexivity)

lemma grad_rec_1_2:
  shows "Cong a b (Gradn a b 1) (Gradn a b 2)"
  by (metis Cong_Gradn_Gradn_Suc Suc_1)

lemma grad_rec_2_3:
  shows "Cong a b (Gradn a b 2) (Gradn a b 3)"
proof (cases "a = b")
  case True
  thus ?thesis 
    using Lem_Gradn_id_n cong_reflexivity by presburger 
next
  case False
  thus ?thesis 
    using Diff__Cong_Gradn_Suc_Gradn_Suc2 numeral_2_eq_2 numeral_3_eq_3 
    by presburger 
qed

lemma grad_rec_a_a:
  shows "(Gradn a a n) = a" 
  by simp

lemma Gradn_uniq_aux_1:
  assumes "A  B" 
  shows "Gradn A B n  Gradn A B (Suc n)" 
proof -
  have "Gradn A B 0  Gradn A B (Suc 0)"
    by (simp add: assms)   
  moreover
  have "n > 0  (Gradn A B n  Gradn A B (Suc n))"
    by (metis Diff__Cong_Gradn_Suc_Gradn_Suc2 assms cong_diff_2 gr0_implies_Suc)  
  ultimately
  show ?thesis
    by blast 
qed

lemma Gradn_uniq_aux_1_aa:
  assumes "A  B" 
  shows "Gradn A B (k + n)  Gradn A B (k + (Suc n))"
proof (induction k)
  case 0
  show "Gradn A B (0 + n)  Gradn A B (0 + (Suc n))"
    using Gradn_uniq_aux_1 assms plus_nat.add_0 by presburger 
next
  case (Suc k)
  show "Gradn A B (Suc k + n)  Gradn A B (Suc k + (Suc n))"
    using Gradn_uniq_aux_1 add_Suc_right assms by presburger 
qed

lemma Gradn_uniq_aux_1_bb:
  assumes "A  B" 
  shows "Gradn A B (k + n)  Gradn A B (k + (Suc (Suc n)))"
proof (induction k)
  case 0
  show "Gradn A B (0 + n)  Gradn A B (0 + (Suc(Suc n)))"
    by (metis Gradn_uniq_aux_1 Diff__Bet_Gradn_Gradn_Suc add.left_neutral 
        assms between_equality_2)
next
  case (Suc k)
  show "Gradn A B ((Suc k) + n)  Gradn A B ((Suc k) + (Suc(Suc n)))"
    by (metis Gradn_uniq_aux_1_aa Diff_Bet_Gradn_Suc_Gradn_Suc2 add_Suc 
        add_Suc_shift assms between_equality_2) 
qed

lemma Gradn_aux_1_0:
  assumes  "A  B" 
  shows "Gradn A B (Suc n)  A"
  by (metis Diff__Bet_Gradn_Suc assms bet_neq32__neq)

lemma Gradn_aux_1_1:
  assumes  "A  B" and
    "n  0"
  shows "Gradn A B (Suc n)  B"
proof -
  obtain m where "n = Suc m"
    using assms(2) not0_implies_Suc by blast
  have "Gradn A B (Suc(Suc m))  B"
  proof (induction m)
    show "Gradn A B (Suc(Suc 0))  B"
      by (metis Gradn_uniq_aux_1 Diff__Bet_Gradn_Suc 
          Diff_Bet_Gradn_Suc_Gradn_Suc2 assms(1) between_equality_2) 
  next
    fix m
    assume "Gradn A B (Suc(Suc m))  B"
    thus "Gradn A B (Suc(Suc(Suc m)))  B"
      by (metis Diff__Bet_Gradn_Suc Diff_Bet_Gradn_Suc_Gradn_Suc2 
          assms(1) between_equality_2) 
  qed
  thus ?thesis
    by (simp add: n = Suc m) 
qed

lemma Gradn_aux_1_1_bis:
  assumes  "A  B" and
    "n  1"
  shows "Gradn A B n  B"
proof (cases "n = 0")
  case True
  thus ?thesis
    using Lem_Gradn_0 assms(1) by presburger
next
  case False 
  then obtain m where "n = Suc m"
    using not0_implies_Suc by presburger 
  hence "m  0"
    using assms(2) by force 
  thus ?thesis 
    using Gradn_aux_1_1 assms(1) n = Suc m by blast 
qed

lemma Gradn_aux_1_2:
  assumes  "A  B" and
    "Gradn A B n = A"
  shows "n = 0" 
proof -
  {
    assume "n  0"
    then obtain m where "n = Suc m" 
      using not0_implies_Suc by presburger
    hence "Gradn A B n  A"
      using Gradn_aux_1_0 assms(1) by blast 
    hence "False"
      using Gradn_aux_1_0 assms(1) assms(2) by blast
  }
  thus ?thesis by blast
qed

lemma Gradn_aux_1_3:
  assumes  "A  B" and
    "Gradn A B n = B"
  shows "n = 1"
  using Gradn_aux_1_1_bis assms(1) assms(2) by blast

lemma Gradn_uniq_aux_2_a:
  assumes "A  B" and
    "n  0"
  shows "Gradn A B 0  Gradn A B n"
  by (metis Gradn_aux_1_2 Lem_Gradn_0 assms(1) assms(2)) 

lemma Gradn_uniq_aux_2:
  assumes "A  B" and
    "n < m"
  shows "Gradn A B n  Gradn A B m" 
proof -
  obtain k where "m = (Suc k) + n"
    by (metis Suc_diff_Suc add.commute add_diff_cancel_left' 
        assms(2) less_imp_add_positive)
  have "Gradn A B n  Gradn A B ((Suc k) + n)"
  proof (induction k)
    case 0
    have "0 + n = n" 
      by simp
    thus "Gradn A B n  Gradn A B ((Suc 0) + n)" 
      using Gradn_uniq_aux_1_aa assms(1) by (metis add.commute) 
  next
    case (Suc k)
    hence "Gradn A B n  Gradn A B ((Suc k) + n)" 
      by blast
    have "Suc ((Suc k) + n) = Suc(Suc(k)) + n"
      by simp
    {
      assume "Gradn A B n = Gradn A B ((Suc (Suc k)) + n)" 
      have "Gradn A B n = Gradn A B ((Suc k) + n)" 
      proof (cases "n = 0")
        case True
        thus ?thesis
          by (metis Gradn_aux_1_2 Lem_Gradn_0 
              Gradn A B n = Gradn A B (Suc (Suc k) + n) add_cancel_left_right 
              assms(1) nat_neq_iff zero_less_Suc) 
      next
        case False
        have "(Suc k) + n = Suc(k + n)" 
          by simp
        hence "Bet A B (Gradn A B ((Suc k)+n))" 
          using assms(1) Diff__Bet_Gradn_Suc by presburger
        hence "Bet A B (Gradn A B n)" 
          using assms(1) Diff__Bet_Gradn_Suc 
            Gradn A B n = Gradn A B (Suc (Suc k) + n) add_Suc by presburger 
        have "Bet A (Gradn A B ((Suc k)+n)) (Gradn A B (Suc((Suc k) +n)))"
          using Diff__Bet_Gradn_Gradn_Suc assms(1) by blast
        hence "Bet A (Gradn A B ((Suc k)+n)) (Gradn A B n)" 
          using Gradn A B n = Gradn A B ((Suc (Suc k)) + n) 
            Suc ((Suc k) + n) = Suc(Suc(k)) + n by simp
        moreover
        have "Bet A (Gradn A B n) (Gradn A B ((Suc k) + n))"
          using Diff_Le__Bet_Gradn_Plus assms(1) by blast 
        ultimately
        show ?thesis
          using between_equality_2 by blast 
      qed
      hence False
        using Suc.IH by blast 
    }
    thus "Gradn A B n  Gradn A B ((Suc (Suc k)) + n)" 
      by blast
  qed
  thus ?thesis
    using m = Suc k + n by blast 
qed

lemma Gradn_uniq:
  assumes "A  B" and
    "Gradn A B n = Gradn A B m"
  shows "n = m" 
proof -
  {
    assume "n  m"
    {
      assume "n < m"
      hence "False"
        using Gradn_uniq_aux_2 assms(1) assms(2) by blast 
    }
    moreover
    {
      assume "m < n"
      hence "False"
        by (metis Gradn_uniq_aux_2 assms(1) assms(2))
    }
    ultimately
    have "False"
      using n  m nat_neq_iff by blast 
  }
  thus ?thesis by blast
qed

lemma Gradn_le_suc_1:
  shows "A (Gradn A B n) Le A (Gradn A B (Suc n))"
  using Bet_Gradn_Gradn_Suc l5_12_a by presburger

lemma Gradn_le_1:
  assumes "m  n"
  shows "A (Gradn A B m) Le A (Gradn A B (Suc n))"
  by (metis Bet_Gradn_Gradn_Suc Lem_Gradn_id_n Diff_Le_Bet__Gradn_Gradn 
      assms bet__le1213 le_Suc_eq)

lemma Gradn_le_suc_2:
  shows "B (Gradn A B (Suc n)) Le B (Gradn A B (Suc(Suc n)))"
  by (metis Bet_Gradn_Gradn_Suc Diff__Bet_Gradn_Suc bet__le1213 
      between_exchange3)

lemma grad_equiv_coq_1:
  shows "Grad A B B"
proof -
  have "(Gradn A B (Suc 0)) = B"
    by auto
  thus ?thesis
    by (metis Grad_def n_not_Suc_n)
qed

lemma grad_aab__ab:
  assumes "Grad A A B" 
  shows "A = B" 
proof -
  obtain n where "B = Gradn A A n"
    using Grad_def assms by blast
  thus ?thesis
    by simp 
qed

lemma grad_stab:
  assumes "Grad A B C" and
    "Bet A C C'" and
    "Cong A B C C'"
  shows "Grad A B C'"
proof (cases "A = B")
  case True
  thus ?thesis
    using assms(1) assms(3) cong_reverse_identity by blast 
next
  case False
  obtain n where "n  0  C = Gradn A B n"
    using Grad_def assms(1) by presburger
  hence "Bet A B (Gradn A B n)"
    using False Diff__Bet_Gradn by blast
  hence "Bet A B C"
    using n  0  C = Gradn A B n by blast 
  hence "C' = Gradn A B (Suc n)"
    using False Diff_Bet_Bet_Cong_Gradn_Suc n  0  C = Gradn A B n 
      assms(2) assms(3) between_exchange4 by blast
  thus ?thesis  
    using Grad_def by blast
qed

lemma grad__bet:
  assumes "Grad A B C"
  shows "Bet A B C"
proof (cases "A = B")
  case True
  thus ?thesis
    by (simp add: between_trivial2)
next
  case False
  obtain n where "n  0  C = Gradn A B n"
    using Grad_def assms(1) by presburger
  hence "Bet A B (Gradn A B n)"
    using False Diff__Bet_Gradn by blast
  thus ?thesis
    using n  0  C = Gradn A B n by blast 
qed

lemma grad__col:
  assumes "Grad A B C"
  shows "Col A B C"
  by (simp add: assms bet_col grad__bet) 

lemma grad_neq__neq13:
  assumes "Grad A B C" and
    "A  B"
  shows "A  C"
  using assms(1) assms(2) between_identity grad__bet by blast

lemma grad_neq__neq12:
  assumes "Grad A B C" and
    "A  C"
  shows "A  B"
  using Grad_def assms(1) assms(2) grad_rec_a_a by force 

lemma grad112__eq:
  assumes "Grad A A B"
  shows "A = B"
  by (meson assms grad_neq__neq12) 

lemma grad121__eq:
  assumes "Grad A B A"
  shows "A = B"
  using assms grad_neq__neq13 by blast

lemma grad__le: 
  assumes "Grad A B C"
  shows "A B Le A C"
  using assms bet__le1213 grad__bet by blast

lemma grad2_init:
  shows "Grad2 A B B C D D"
proof -
  have "(B = Gradn A B (Suc 0))  (D = Gradn C D (Suc 0))"
    using One_nat_def Lem_Gradn_1 by presburger 
  thus ?thesis
    using Grad2_def by blast
qed

lemma Grad2_stab:
  assumes "Grad2 A B C D E F" and
    "Bet A C C'" and
    "Cong A B C C'" and
    "Bet D F F'" and
    "Cong D E F F'"
  shows "Grad2 A B C' D E F'"
proof -
  obtain n where "(n  0)  (C = Gradn A B n)  (F = Gradn D E n)"
    using Grad2_def assms(1) by presburger
  have "C' = Gradn A B (Suc n)"
    by (metis Diff_Bet_Bet_Cong_Gradn_Suc Lem_Gradn_id_n Diff__Bet_Gradn
        n  0  C = Gradn A B n  F = Gradn D E n assms(2) assms(3) 
        between_exchange4 cong_reverse_identity)
  moreover
  have "F' = Gradn D E (Suc n)"
    by (metis Diff_Bet_Bet_Cong_Gradn_Suc Lem_Gradn_id_n Diff__Bet_Gradn
        n  0  C = Gradn A B n  F = Gradn D E n assms(4) assms(5) 
        between_exchange4 cong_reverse_identity)
  ultimately
  show ?thesis
    using Grad2_def by blast
qed

lemma bet_cong2_grad__grad2_aux_1:
  assumes "C = (Gradn A B 0)" and
    "Bet D E F" and
    "Cong A B D E" and
    "Cong B C E F"
  shows "F = Gradn D E 2"
proof -
  have "(Gradn A B 0) = A"
    using Lem_Gradn_0 by blast
  hence "A = C" 
    using assms(1) by auto
  hence "Cong D E E F"
    using assms(3) assms(4) cong_transitivity not_cong_4312 by blast
  thus ?thesis
    by (metis Diff_Bet_Bet_Cong_Gradn_Suc Lem_Gradn_1 Suc_1 
        assms(2) cong_diff_3 grad_rec_1_2)
qed

lemma bet_cong2_grad__grad2_aux_2:
  assumes "Bet D E F" and
    "Cong A B D E" and
    "Cong B (Gradn A B (Suc n)) E F"
  shows "F = Gradn D E (Suc n)"
proof -
  have " A B D E F. (Bet D E F  Cong A B D E  
  Cong B (Gradn A B (Suc n)) E F  F = Gradn D E (Suc n))"
  proof (induction n)
    case 0
    { 
      fix A B C D E F
      assume 1: "Bet D E F  Cong A B D E  Cong B (Gradn A B (Suc 0)) E F"
      hence "Gradn A B (Suc 0) = B"
        using One_nat_def Lem_Gradn_1 by presburger
      hence "E = F"
        by (metis "1" cong_diff_4)
      hence "F = Gradn D E (Suc 0)"
        by simp 
    }
    thus ?case
      by blast
  next
    case (Suc n)
    {
      assume 2: " A B D E F. (Bet D E F  Cong A B D E  
    Cong B (Gradn A B (Suc n)) E F)  F = Gradn D E (Suc n)"
      {
        fix A B D E F
        assume "Bet D E F" and 
          "Cong A B D E" and 
          "Cong B (Gradn A B (Suc (Suc n))) E F"
        have "Cong A B (Gradn A B (Suc n)) (Gradn A B (Suc(Suc n)))"
          using Cong_Gradn_Suc_Gradn_Suc2 by auto
        have "Bet A (Gradn A B (Suc n)) (Gradn A B (Suc(Suc n)))"
          using Bet_Gradn_Gradn_Suc by auto
        have "F = Gradn D E (Suc (Suc n))" 
        proof (cases "A = B")
          case True
          thus ?thesis 
            by (metis Cong A B D E Cong B (Gradn A B (Suc (Suc n))) E F 
                cong_reverse_identity grad_rec_a_a)
        next
          case False
          have "D  E"
            using False Cong A B D E cong_diff by blast 
          obtain F' where "Bet D E F'" and "Cong E F' B (Gradn A B (Suc n))"
            using segment_construction by fastforce
          have "Cong B (Gradn A B (Suc n)) E F'"
            using Cong E F' B (Gradn A B (Suc n)) not_cong_3412 by blast 
          hence "F' = Gradn D E (Suc n)"
            using Bet D E F' Cong A B D E "2" by blast
          thus ?thesis 
          proof (cases "E = F'")
            case True
            thus ?thesis 
              by (metis Diff_Bet_Bet_Cong_Gradn_Suc
                  Cong A B (Gradn A B (Suc n)) (Gradn A B (Suc (Suc n))) D  E 
                  F' = Gradn D E (Suc n) cong_inner_transitivity 
                  cong_reverse_identity Bet D E F Cong A B D E 
                  Cong B (Gradn A B (Suc (Suc n))) E F 
                  Cong E F' B (Gradn A B (Suc n)))
          next
            case False
            have "F = Sym D E (Gradn D E ((Suc(Suc n))-1))" 
            proof -
              have "Bet D E (Gradn D E ((Suc(Suc n))-1))"
                using Bet D E F' F' = Gradn D E (Suc n) by auto 
              moreover
              have "PreGrad D E (Gradn D E ((Suc(Suc n))-1)) F" 
              proof -
                have "Bet D E (Gradn D E (Suc n))"
                  using Diff__Bet_Gradn_Suc D  E by blast 
                hence "Bet D E (Gradn D E ((Suc(Suc n))-1))"
                  by fastforce 
                moreover
                have "Bet D (Gradn D E (Suc n)) F" 
                proof -
                  have "Bet B (Gradn A B (Suc n)) (Gradn A B (Suc(Suc n)))"
                    by (metis Diff__Bet_Gradn_Suc 
                        Bet A (Gradn A B (Suc n)) (Gradn A B (Suc (Suc n))) 
                        between_exchange3) 
                  moreover
                  have "Cong B (Gradn A B (Suc n)) E (Gradn D E (Suc n))"
                    using Cong B (Gradn A B (Suc n)) E F' 
                      F' = Gradn D E (Suc n) by blast 
                  moreover
                  have "Cong B (Gradn A B (Suc(Suc n))) E (Gradn D E (Suc(Suc n)))"
                  proof -
                    have "Bet A B (Gradn A B (Suc n))" 
                      using Diff__Bet_Gradn_Suc not_bet_distincts by blast
                    moreover
                    have "Cong D E (Gradn D E (Suc n)) (Gradn D E (Suc (Suc n)))" 
                      using Cong_Gradn_Suc_Gradn_Suc2 by auto
                    moreover
                    have "Bet A (Gradn A B (Suc n)) (Gradn A B (Suc (Suc n)))" 
                      using Bet_Gradn_Gradn_Suc by blast 
                    moreover
                    have "Cong D E (Gradn A B (Suc n)) (Gradn A B (Suc (Suc n)))"                      
                      using Cong A B (Gradn A B (Suc n)) (Gradn A B (Suc (Suc n))) 
                        Cong A B D E cong_inner_transitivity by blast
                    moreover
                    have "Bet D (Gradn D E (Suc n)) (Gradn D E (Suc (Suc n)))" 
                      using Bet_Gradn_Gradn_Suc by blast
                    thus ?thesis
                      using l2_11_b between_exchange3 calculation(2) cong_inner_transitivity 
                      by (meson Bet B (Gradn A B (Suc n)) (Gradn A B (Suc (Suc n))) 
                          Cong B (Gradn A B (Suc n)) E (Gradn D E (Suc n)) 
                          Bet D E (Gradn D E (Suc n))
                          calculation(3) calculation(4))
                  qed
                  moreover
                  have "E Out (Gradn D E (Suc(Suc n))) (Gradn D E (Suc n))"
                    by (metis False Diff__Bet_Gradn_Gradn_Suc 
                        Bet D E (Gradn D E (Suc n)) 
                        D  E F' = Gradn D E (Suc n) 
                        bet_out between_exchange3 l6_6) 
                  ultimately
                  show ?thesis
                    using Bet_Gradn_Gradn_Suc Diff__Bet_Gradn_Suc D  E 
                      construction_uniqueness not_cong_3412 
                    by (metis Bet D E F Cong B (Gradn A B (Suc (Suc n))) E F)
                qed
                hence "Bet D (Gradn D E ((Suc(Suc n))-1)) F"
                  by simp 
                moreover
                have "Cong (Gradn A B (Suc n)) (Gradn A B(Suc(Suc n))) (Gradn D E (Suc n)) F"
                proof -
                  let ?A = "A"
                  let ?B = "Gradn A B (Suc n)"
                  let ?C = "Gradn A B (Suc (Suc n))"
                  let ?A' = "D"
                  let ?B' = "Gradn D E (Suc n)"
                  let ?C' = "F"
                  have "Bet ?A ?B ?C" 
                    using Bet A (Gradn A B (Suc n)) (Gradn A B (Suc (Suc n))) by auto
                  moreover have "Bet ?A' ?B' ?C'" 
                    using Bet D (Gradn D E (Suc n)) F by auto
                  moreover have "Cong ?A ?B ?A' ?B'" 
                    by (metis Diff__Bet_Gradn_Suc cong_reverse_identity
                        Cong A B D E Cong E F' B (Gradn A B (Suc n)) 
                        F' = Gradn D E (Suc n) l2_11_b not_cong_3412)
                  moreover have "Cong ?A ?C ?A' ?C'" 
                    by (metis Diff__Bet_Gradn_Suc cong_reverse_identity 
                        Bet D E F Cong A B D E 
                        Cong B (Gradn A B (Suc (Suc n))) E F l2_11_b)
                  ultimately show ?thesis 
                    using l4_3_1 by blast
                qed
                hence "Cong A B (Gradn D E (Suc n)) F"
                  using Cong A B (Gradn A B (Suc n)) (Gradn A B (Suc (Suc n))) 
                    cong_transitivity by blast 
                hence "Cong D E (Gradn D E (Suc n)) F" 
                  using Cong A B D E cong_inner_transitivity by blast
                hence "Cong D E (Gradn D E ((Suc(Suc n))-1)) F" 
                  by simp
                ultimately
                show ?thesis
                  using D  E PreGrad_def by blast 
              qed
              ultimately
              show ?thesis
                by (metis PreGrad_def LemSym_aux) 
            qed
            hence "F = Gradn D E (Suc (Suc n))"
              by simp 
            thus ?thesis by blast
          qed
        qed
      }
      hence " A B D E F. (Bet D E F  Cong A B D E  
    Cong B (Gradn A B (Suc(Suc n))) E F)  F = Gradn D E (Suc(Suc n))" 
        by blast
    }
    thus ?case
      using Suc.IH by fastforce 
  qed
  thus ?thesis
    using assms(1) assms(2) assms(3) by blast 
qed

lemma bet_cong2_grad__grad2_aux:
  assumes "n  0" and
    "C = (Gradn A B n)" and
    "Bet D E F" and
    "Cong A B D E" and
    "Cong B C E F"
  shows "F = Gradn D E n"
proof -
  obtain k where "n = Suc k"
    using assms(1) not0_implies_Suc by blast
  thus ?thesis
    using assms(5) assms(2) assms (3) assms(4) bet_cong2_grad__grad2_aux_2 by blast
qed

lemma bet_cong2_grad__grad2:
  assumes "Grad A B C" and
    "Bet D E F" and
    "Cong A B D E" and
    "Cong B C E F"
  shows "Grad2 A B C D E F" 
proof (cases "A = B")
  case True
  thus ?thesis 
    by (metis assms(1) assms(4) cong_diff_4 grad2_init grad_neq__neq12)
next
  case False
  hence "D  E"
    using assms(3) cong_diff by blast
  obtain n where "n  0  C = (Gradn A B n)"
    using Grad_def assms(1) by presburger 
  have "F = Gradn D E n" 
    using bet_cong2_grad__grad2_aux n  0  C = Gradn A B n 
      assms(2) assms(3) assms(4) by blast 
  thus ?thesis 
    using n  0  C = (Gradn A B n) Grad2_def by blast 
qed

lemma grad2__grad123: 
  assumes "Grad2 A B C D E F"
  shows "Grad A B C" 
proof -
  obtain n where "(n  0)  (C = Gradn A B n)  (F = Gradn D E n)"
    using Grad2_def assms by presburger
  thus ?thesis
    using Grad_def by blast 
qed

lemma grad2__grad456:
  assumes "Grad2 A B C D E F"
  shows "Grad D E F" 
proof -
  obtain n where "(n  0)  (C = Gradn A B n)  (F = Gradn D E n)"
    using Grad2_def assms by presburger
  thus ?thesis
    using Grad_def by blast 
qed

lemma grad_sum_aux_R1: 
  assumes 
    "A C Le A D" and
    "Cong A D A E" and
    "Cong A C A E'" and 
    "A Out E E'"
  shows "Bet A E' E"
  by (meson Out_cases l5_6 assms(2) assms(3) assms(4) assms(1) l6_13_1) 

lemma grad_sum_aux_0: 
  assumes "A  B" and
    "D = Gradn A B (Suc(Suc n))" and
    "Cong A D A E" and
    "C = Gradn A B (Suc n)" and
    "Cong A C A E'" and 
    "A Out E E'"
  shows "Bet A E' E" 
proof -
  {
    assume "Bet A E E'"
    have False
      by (metis Gradn_aux_1_0 Gradn_uniq_aux_1 
          Diff__Bet_Gradn_Gradn_Suc Bet A E E' 
          assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) 
          between_cong_2 between_equality_2 cong_preserves_bet 
          cong_transitivity l5_1 l6_6 not_cong_3412)
  }
  thus ?thesis
    using Out_def assms(6) by auto 
qed


lemma grad_sum_aux_1:
  assumes "A  B" and
    "D = Gradn A B (Suc(Suc n))" and
    "Bet A C E" and
    "Cong A D C E" and
    "F = Gradn A B (Suc n)" and
    "Bet A C E'" and 
    "Cong A F C E'" and 
    "C  A" 
  shows "Bet A E' E" 
proof -
  have "A  B  D = Gradn A B (Suc(Suc n)) 
  Bet A C E  Cong A D C E 
  F = Gradn A B (Suc n) 
  Bet A C E'  Cong A F C E'  A  C 
  Bet A E' E"
  proof (induction n)
    case 0
    {
      fix A B C D E F
      assume "A  B" and
        "D = Gradn A B (Suc(Suc 0))" and
        "Bet A C E" and
        "Cong A D C E" and
        "F = Gradn A B (Suc 0)" and
        "Bet A C E'" and
        "Cong A F C E'" and
        "A  C"
      have "F = B"
        by (simp add: F = Gradn A B (Suc 0))
      have "Bet C E' E" 
      proof -
        have "C Out E' E" 
        proof -
          {
            assume "C = E'"
            hence "A = F"
              using Cong A F C E' cong_identity by blast 
            hence False
              using A  B F = B by fastforce
          }
          hence "C  E'"
            by  blast
          moreover
          {
            assume "C = E"
            hence "A = D"
              using Cong A D C E cong_identity by blast 
            hence False
              by (metis Gradn_aux_1_0 A  B D = Gradn A B (Suc (Suc 0))) 
          }
          hence "C  E" 
            by blast
          moreover
          have "Bet E' C A" 
            using Bet A C E' Bet_cases by blast
          moreover
          have "Bet E C A"
            using Bet_cases Bet A C E by blast 
          ultimately
          show ?thesis
            using A  C l6_2 by metis 
        qed
        moreover
        have "A F Le A D" 
          using F = Gradn A B (Suc 0) D = Gradn A B (Suc(Suc 0)) 
            Gradn_le_suc_1 by blast 
        hence "C E' Le C E" 
          using l5_6 Cong A D C E Cong A F C E' by blast 
        ultimately
        show ?thesis 
          using l6_13_1 by blast
      qed
      hence "Bet A E' E"
        by (meson Bet A C E between_exchange2) 
    }
    thus ?case 
      by force
  next
    case (Suc n)
    {
      assume 1: "A  B  D = Gradn A B (Suc(Suc n)) 
    Bet A C E  Cong A D C E 
    F = Gradn A B (Suc n) 
    Bet A C E'  Cong A F C E'  A  C 
    Bet A E' E" 
      {
        assume "A  B" and
          "D = Gradn A B (Suc(Suc(Suc n)))" and
          "Bet A C E" and
          "Cong A D C E" and
          "F = Gradn A B (Suc(Suc n))" and
          "Bet A C E'"
          "Cong A F C E'" and
          "A  C"
        obtain F' where "F' = Gradn A B (Suc n)"
          by auto
        obtain E'' where "Bet A C E''  Cong A F' C E''"
          by (meson Cong_cases segment_construction)
        have "Bet A E'' E'"
        proof -
          have "A  B"
            by (simp add: assms(1)) 
          moreover
          have "F = Gradn A B (Suc(Suc n))"
            by (simp add: F = Gradn A B (Suc (Suc n))) 
          moreover
          have "Bet A C E'"
            by (simp add: assms(6)) 
          moreover
          have "Cong A F C E'"
            by (simp add: assms(7)) 
          moreover
          have "F' = Gradn A B (Suc n)"
            by (simp add: F' = Gradn A B (Suc n)) 
          moreover
          have "Bet A C E''"
            by (simp add: Bet A C E''  Cong A F' C E'') 
          moreover
          have "Cong A F' C E''"
            by (simp add: Bet A C E''  Cong A F' C E'')
          moreover
          have "A  C"
            using assms(8) by auto
          ultimately
          show ?thesis 
          proof -
            {
              assume "Bet C E'' E'"
              hence ?thesis 
                using assms(6) between_exchange2 by blast
            }
            moreover
            {
              assume "Bet C E' E''"
              hence ?thesis
                by (metis Diff__Bet_Gradn_Gradn_Suc Gradn_aux_1_0 
                    bet2__out cong_identity_inv cong_preserves_bet 
                    Bet A C E''  Cong A F' C E'' F = Gradn A B (Suc (Suc n)) 
                    F' = Gradn A B (Suc n) assms(1) assms(7) 
                    calculation not_bet_distincts)
            }
            ultimately show ?thesis 
              by (metis Bet A C E'' assms(6) assms(8) l5_2)
          qed
        qed
        have "C E'' Le C E'"
          using Bet A C E''  Cong A F' C E'' Bet A E'' E' 
            bet__le1213 between_exchange3 by blast 
        have "A (Gradn A B (Suc (Suc n))) Le A (Gradn A B (Suc (Suc (Suc n))))"
          using Gradn_Le_Gradn_Suc by blast 
        hence "A F Le A D"
          using F = Gradn A B (Suc (Suc n)) 
            D = Gradn A B (Suc(Suc(Suc n))) by blast
        hence "C E' Le C E"
          using assms(4) assms(7) l5_6 by blast 
        have "Bet A E' E"
        proof -
          {
            assume "Bet A E E'"
            hence "A E Le A E'"
              using bet__le1213 by auto
            hence "C E Le C E'"
              using Bet A E E' assms(3) bet__le1213 between_exchange3 by blast
            hence "Cong C E C E'"
              by (simp add: C E' Le C E le_anti_symmetry)
            hence "Cong A D A F"
              by (meson assms(4) assms(7) cong_transitivity not_cong_3412) 
            hence False
              by (metis Gradn_uniq_aux_1 Diff__Bet_Gradn_Gradn_Suc 
                  D = Gradn A B (Suc (Suc (Suc n))) 
                  F = Gradn A B (Suc (Suc n)) 
                  assms(1) between_cong not_cong_3412) 
          }
          thus ?thesis
            by (metis assms(3) assms(6) assms(8) l5_1) 
        qed
      }
    }
    thus ?case    
      using Suc.IH by fastforce 
  qed
  thus ?thesis
    using assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) 
      assms(7) assms(8) by blast 
qed

lemma grad_sum_aux_1A:
  assumes "A  B" and
    "C = Gradn A B (Suc (Suc 0))"
  shows "B Midpoint A C" 
proof -
  have "C =  (Sym A B (Gradn A B (Suc 0)))"
    by (simp add: assms(2))
  have "... = (Sym A B B)"
    by (simp add: assms(1)) 
  hence "(SOME x::'p. PreGrad A B B x) = C"
    using assms(1) by (metis Gradn_aux_1_0 Sym.simps 
        C = Sym A B (Gradn A B (Suc 0)) assms(2)) 
  hence "PreGrad A B B C"
    by (metis PreGrad_def Sym_Bet__Bet_Bet Sym_Bet__Cong 
        Sym.elims assms(1) not_bet_distincts)
  thus ?thesis
    by (simp add: Midpoint_def PreGrad_def)
qed

lemma grad_sum_aux_2:
  assumes "A  B" and
    "D = Gradn A B (Suc(Suc n))" and
    "Bet A C E" and
    "Cong A D C E" and
    "F = Gradn A B (Suc n)" and
    "Bet A C E'" and 
    "Cong A F C E'" and 
    "C  A"
  shows "Cong A B E' E" 
proof -
  have "A  B  D = Gradn A B (Suc(Suc n))  Bet A C E  Cong A D C E 
  F = Gradn A B (Suc n)  Bet A C E'  Cong A F C E'  A  C 
  Cong  A B E' E"
  proof (induction n)
    case 0
    {
      fix A B C D E F
      assume "A  B" and
        "D = Gradn A B (Suc(Suc 0))" and
        "Bet A C E" and
        "Cong A D C E" and
        "F = Gradn A B (Suc 0)" and
        "Bet A C E'" and
        "Cong A F C E'" and
        "A  C"
      have "B Midpoint A D" 
        using grad_sum_aux_1A A  B D = Gradn A B (Suc (Suc 0)) by blast 
      hence "Bet A B D  Cong A B B D"
        using Midpoint_def by blast
      have "B = F"
        using One_nat_def Lem_Gradn_1 F = Gradn A B (Suc 0) by presburger 
      hence "Cong A B C E'"
        by (simp add: Cong A F C E') 
      have "Cong B D C E'"
        using B = F Bet A B D  Cong A B B D Cong A B C E' 
          cong_inner_transitivity by blast 
      have "Cong B D E' E" 
        using cong_commutativity cong_diff_2 Tarski_neutral_dimensionless_axioms 
          A  B A  C B = F Bet A B D  Cong A B B D Bet A C E' 
          Bet A C E Cong A D C E Cong A F C E' bet_out between_symmetry 
          l6_2 out_cong_cong by metis
      hence "Cong C E' E' E"
        using Cong B D C E' cong_inner_transitivity by blast 
      hence "Cong A B E' E"
        using B = F Cong A F C E' cong_transitivity by blast 
    }
    thus ?case
      by blast 
  next
    case (Suc n)
    have "Cong A B E' E" 
    proof -
      have "Bet A F D" 
        using Diff__Bet_Gradn_Gradn_Suc assms(1) assms(2) assms(5) by blast
      have "Cong A B F D" 
        using  Diff__Cong_Gradn_Suc_Gradn_Suc2 assms(1) assms(2) 
          assms(5) by blast
      have "Bet A E' E" 
        using grad_sum_aux_1 assms(1) assms(2) assms(3) assms(4) 
          assms(5) assms(6) assms(7) assms(8) by blast
      {
        assume "Bet C E E'"
        hence "Cong A B E' E" 
          by (metis Bet A E' E Bet A F D Cong A B F D 
              assms(1) assms(4) assms(6) assms(7) between_cong between_equality_2 
              between_exchange3 cong_diff cong_transitivity not_cong_3412) 
      }
      moreover
      {
        assume "Bet C E' E"
        hence "Cong A B E' E"  
          using Bet A F D Cong A B F D assms(4) assms(7) 
            cong_transitivity l4_3_1 by blast
      }
      ultimately show ?thesis 
        by (metis assms(3) assms(6) assms(8) l5_2)
    qed
    thus ?case 
      by blast
  qed
  thus ?thesis 
    using assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) assms(8) by blast
qed

lemma grad_sum_aux:
  assumes "A  B" and
    "C = Gradn A B (Suc n)" and
    "D = Gradn A B (Suc m)" and
    "Bet A C E" and
    "Cong A D C E"
  shows "E = Gradn A B ((Suc n) + (Suc m))" 
proof -
  have " A B C D E. (A  B  C = Gradn A B (Suc n)  
  D = Gradn A B (Suc m)  Bet A C E  Cong A D C E) 
  E = Gradn A B ((Suc n) + (Suc m))"
  proof (induction m)
    case 0
    {
      fix A B C D E
      assume "A  B" and
        "C = Gradn A B (Suc n)" and
        "D = Gradn A B (Suc 0)" and  
        "Bet A C E" and
        "Cong A D C E"
      have "E = Gradn A B ((Suc n) + (Suc 0))"
        by (metis Nat.add_0_right Diff_Bet_Bet_Cong_Gradn_Suc 
            Lem_Gradn_0 Diff__Bet_Gradn_Suc Cong_Gradn_Gradn_Suc 
            cong_transitivity A  B Bet A C E C = Gradn A B (Suc n) 
            Cong A D C E D = Gradn A B (Suc 0) add_Suc_shift between_exchange4)
    }
    thus ?case
      by force 
  next
    case (Suc m)
    {
      assume 1: " A B C D E. (A  B  C = Gradn A B (Suc n)  
    D = Gradn A B (Suc m)  Bet A C E  Cong A D C E) 
    E = Gradn A B ((Suc n) + (Suc m))"
      {
        fix A B C D E
        assume "A  B" and
          "C = Gradn A B (Suc n)" and
          "D = Gradn A B (Suc(Suc m))" and
          "Bet A C E" and
          "Cong A D C E"
        obtain F where 2: "F = Gradn A B (Suc m)"
          by simp
        obtain E' where 3: "Bet A C E'  Cong A F C E'"
          by (meson cong_4312 segment_construction)
        hence 4: "E' = Gradn A B ((Suc n) + (Suc m))"
          using 1 2 A  B C = Gradn A B (Suc n) by blast
        have 5: "Bet A B E" 
        proof -
          have "Grad A B C" 
            using Grad_def C = Gradn A B (Suc n) by blast 
          hence "Bet A B C"
            using grad__bet by blast 
          thus ?thesis
            using Bet A C E between_exchange4 by blast 
        qed
        have 6: "Bet A E' E" 
          using grad_sum_aux_1 A  B D = Gradn A B (Suc(Suc m)) 
            Bet A C E Cong A D C E 2 3 Gradn_aux_1_0 
            C = Gradn A B (Suc n) by blast 
        have "Cong A B E' E"
          using grad_sum_aux_2 "2" "3" Gradn_aux_1_0 A  B 
            Bet A C E C = Gradn A B (Suc n) Cong A D C E 
            D = Gradn A B (Suc (Suc m)) by blast
        hence "E = Gradn A B (Suc(Suc n) + (Suc m))" 
          using Diff_Bet_Bet_Cong_Gradn_Suc 3 4 5 6
          by (metis A  B add_Suc_right add_Suc_shift) 
        hence "E = Gradn A B ((Suc n) + (Suc(Suc m)))"
          by simp 
      }
      hence " A B C D E. (A  B  C = Gradn A B (Suc n)  
    D = Gradn A B (Suc(Suc m))  Bet A C E  Cong A D C E) 
    E = Gradn A B ((Suc n) + (Suc(Suc m)))" by blast
    }
    thus ?case
      using Suc.IH by fastforce 
  qed
  thus ?thesis
    using assms(1) assms(2) assms(3) assms(4) assms(5) by blast 
qed

lemma grad_sum:
  assumes "Grad A B C" and
    "Grad A B D" and
    "Bet A C E" and
    "Cong A D C E"
  shows "Grad A B E"
proof (cases "A = B")
  case True
  thus ?thesis
    using assms(1) assms(2) assms(3) assms(4) grad112__eq grad_stab by blast 
next
  case False
  obtain n where 1: "(n  0)  (C = Gradn A B n)"
    using Grad_def assms(1) by presburger 
  obtain m where 2: "(m  0)  (D = Gradn A B m)"
    using Grad_def assms(2) by presburger 
  obtain k l where "n = Suc k  m = Suc l" 
    using 1 2 not0_implies_Suc by presburger 
  hence "E = Gradn A B (n + m)" 
    using False 1 2 grad_sum_aux assms(3) assms(4) by blast 
  thus ?thesis
    by (meson Grad_def n  0  C = Gradn A B n add_is_0) 
qed

lemma SymR_ex:
  assumes "B Midpoint A C"
  shows "C = SymR A B"
proof -
  have " x. B Midpoint A x" 
    using assms(1) by blast
  thus ?thesis
    using assms(1) someI_ex
    by (metis SymR_uniq_aux SymR.elims) 
qed

lemma SymR__midp:
  assumes "C = SymR A B"
  shows "B Midpoint A C"
  using SymR_ex assms symmetric_point_construction by blast

lemma SymR_uniq:
  assumes "C = SymR A B" and
    "D = SymR A B" 
  shows "C = D"
  by (simp add: assms(1) assms(2))

lemma GradExpn_1:
  shows "GradExpn A A n = A" 
  by simp

lemma GradExpn_2:
  shows "Bet A B (GradExpn A B (Suc n))"
proof (cases "A = B")
  case True
  thus ?thesis
    using between_trivial2 by blast
next
  case False
  show "Bet A B (GradExpn A B (Suc n))" 
  proof(induction n)
    case 0
    have "(GradExpn A B 1) = B" 
      using False by simp
    thus "Bet A B (GradExpn A B (Suc 0))"
      by (metis One_nat_def not_bet_distincts) 
  next
    case (Suc n)
    {
      assume "Bet A B (GradExpn A B (Suc n))"
      have "(Suc (Suc n))  0  (Suc (Suc n))  1" 
        by presburger
      hence "GradExpn A B (Suc(Suc n)) = SymR A (GradExpn A B (Suc n))" 
        using False by simp             
      obtain C where "(GradExpn A B (Suc n)) Midpoint A C"
        using symmetric_point_construction by blast
      hence "Bet A (GradExpn A B (Suc n)) C" 
        using Midpoint_def by blast
      have "C = GradExpn A B (Suc(Suc n))"
        using SymR_ex 
          GradExpn A B (Suc (Suc n)) = SymR A (GradExpn A B (Suc n)) 
          GradExpn A B (Suc n) Midpoint A C by presburger 
      hence "Bet A B (GradExpn A B (Suc(Suc n)))"
        using Bet A (GradExpn A B (Suc n)) C 
          Bet A B (GradExpn A B (Suc n)) between_exchange4 by blast
    }
    thus ?case 
      using Suc.IH by blast 
  qed
qed

lemma GradExpn_3:
  shows "Bet A (GradExpn A B (Suc n)) (GradExpn A B (Suc(Suc n)))"
proof (cases "A = B")
  case True
  thus ?thesis
    using between_trivial2 by force
next
  case False
  show "Bet A (GradExpn A B (Suc n)) (GradExpn A B (Suc(Suc n)))" 
  proof(induction n)
    case 0
    have "(GradExpn A B 1) = B" 
      using False by simp
    thus "Bet A (GradExpn A B (Suc 0)) (GradExpn A B (Suc(Suc 0)))"
      using GradExpn_2 One_nat_def by presburger
  next
    case (Suc n)
    {
      assume "Bet A (GradExpn A B (Suc n)) (GradExpn A B (Suc(Suc n)))"
      have "(Suc(Suc (Suc n)))  0  (Suc(Suc (Suc n)))  1" 
        by presburger
      hence "GradExpn A B (Suc(Suc(Suc n))) = SymR A (GradExpn A B (Suc(Suc n)))" 
        using False by simp             
      obtain C where "(GradExpn A B (Suc(Suc n))) Midpoint A C"
        using symmetric_point_construction by blast
      hence "Bet A (GradExpn A B (Suc(Suc n))) C" 
        using Midpoint_def by blast
      have "C = GradExpn A B (Suc(Suc(Suc n)))"
        using SymR_ex 
          GradExpn A B (Suc(Suc (Suc n))) = SymR A (GradExpn A B (Suc(Suc n))) 
          GradExpn A B (Suc(Suc n)) Midpoint A C by presburger 
      hence "Bet A (GradExpn A B (Suc (Suc(n)))) (GradExpn A B (Suc(Suc(Suc n))))"
        using Bet A (GradExpn A B (Suc (Suc n))) C by blast 
    }
    thus ?case 
      using Suc.IH by blast 
  qed
qed

lemma GradExpn_4:
  shows "Cong A (GradExpn A B (Suc n)) (GradExpn A B (Suc n)) (GradExpn A B (Suc(Suc n)))"
proof (cases "A = B")
  case True
  thus ?thesis
    using GradExpn_1 cong_trivial_identity by presburger
next
  case False
  show "Cong A (GradExpn A B (Suc n)) (GradExpn A B (Suc n)) (GradExpn A B (Suc(Suc n)))"
  proof (induction n)
    case 0
    have "GradExpn A B (Suc 0) = B"
      using False by simp
    show "Cong A (GradExpn A B (Suc 0)) (GradExpn A B (Suc 0)) (GradExpn A B (Suc(Suc 0)))" 
    proof -
      have "(Suc(Suc 0))  0  (Suc(Suc 0))  1" 
        by presburger
      hence "GradExpn A B (Suc(Suc 0)) = SymR A (GradExpn A B (Suc 0))" 
        using False by simp             
      obtain C where "(GradExpn A B (Suc 0)) Midpoint A C"
        using symmetric_point_construction by blast
      hence "Cong A (GradExpn A B (Suc 0)) (GradExpn A B (Suc 0)) C" 
        using Midpoint_def by blast
      have "C = GradExpn A B (Suc(Suc 0))"
        using SymR_ex GradExpn A B (Suc (Suc 0)) = SymR A (GradExpn A B (Suc 0)) 
          GradExpn A B (Suc 0) Midpoint A C by presburger 
      thus ?thesis
        using Cong A (GradExpn A B (Suc 0)) (GradExpn A B (Suc 0)) C by blast 
    qed
  next
    case (Suc n)
    {
      assume "Cong A (GradExpn A B (Suc n)) (GradExpn A B (Suc n)) (GradExpn A B (Suc(Suc n)))"
      have "(Suc(Suc(Suc n)))  0  (Suc(Suc(Suc n)))  1" 
        by presburger
      hence "GradExpn A B (Suc(Suc(Suc n))) = SymR A (GradExpn A B (Suc(Suc n)))" 
        using False by simp             
      obtain C where "(GradExpn A B (Suc(Suc n))) Midpoint A C"
        using symmetric_point_construction by blast
      hence "Cong A (GradExpn A B (Suc(Suc n))) (GradExpn A B (Suc(Suc n))) C" 
        using Midpoint_def by blast
      have "C = GradExpn A B (Suc(Suc(Suc n)))"
        using SymR_ex 
          GradExpn A B (Suc(Suc (Suc n))) = SymR A (GradExpn A B (Suc(Suc n))) 
          GradExpn A B (Suc(Suc n)) Midpoint A C by presburger 
      hence "Cong A (GradExpn A B (Suc(Suc n)))
    (GradExpn A B (Suc(Suc n))) (GradExpn A B (Suc(Suc(Suc n))))"
        using Cong A (GradExpn A B (Suc (Suc n))) (GradExpn A B (Suc (Suc n))) C 
        by fastforce 
    }
    thus ?case
      using Suc.IH by blast
  qed
qed

lemma gradexp_init:
  shows "GradExp A B B"
proof -
  have "B = GradExpn A B (Suc 0)"
    by auto
  thus ?thesis
    using GradExp_def by blast
qed

lemma gradexp_stab_aux_0:
  assumes "C = GradExpn A B 0" and
    "Bet A C C'" and
    "Cong A C C C'"
  shows "C' = GradExpn A B 0" 
proof -
  obtain m where "m = Suc 0"
    by simp 
  show ?thesis
  proof (cases "A = B")
    case True
    thus ?thesis
      by (metis GradExpn_1 assms(1) assms(3) cong_reverse_identity)
  next
    case False
    hence "A  B" 
      by auto
    show ?thesis
    proof -
      have "C = A" 
        using assms(1) by simp
      hence "C = C'"
        using assms(3) cong_reverse_identity by blast 
      thus ?thesis
        using assms(1) by fastforce 
    qed
  qed
qed

lemma gradexp_stab_aux_n:
  assumes "C = GradExpn A B (Suc n)" and
    "Bet A C C'" and
    "Cong A C C C'"
  shows "C' = GradExpn A B (Suc(Suc n))" 
proof -
  obtain m where "m = Suc(Suc n)"
    by simp 
  hence "m - 1 = Suc n" 
    by simp
  show ?thesis
  proof (cases "A = B")
    case True
    thus ?thesis
      by (metis GradExpn_1 assms(1) assms(3) cong_reverse_identity)
  next
    case False
    hence "A  B" 
      by auto
    show ?thesis
    proof -
      have "C Midpoint A C'"
        using assms(2) assms(3) midpoint_def by auto 
      hence "C' = SymR A C"
        using SymR_ex by blast
      hence "C' = SymR A (GradExpn A B (m - 1))"
        using m - 1 = Suc n assms(1) by presburger 
      hence "C' = GradExpn A B m"
        by (simp add: False m = Suc (Suc n)) 
      thus ?thesis
        using m = Suc (Suc n) by blast
    qed
  qed
qed

lemma gradexp_stab:
  assumes "GradExp A B C" and
    "Bet A C C'" and
    "Cong A C C C'"
  shows "GradExp A B C'" 
proof -
  obtain n where "C = GradExpn A B n"
    using GradExp_def assms(1) by blast 
  show ?thesis 
  proof (cases "n = 0")
    case True
    thus ?thesis
      using C = GradExpn A B n assms(1) assms(2) assms(3) 
        gradexp_stab_aux_0 by blast 
  next
    case False
    then obtain m where "Suc m = n"
      using not0_implies_Suc by auto
    hence "C = GradExpn A B (Suc m)"
      by (simp add: C = GradExpn A B n)
    hence "C' = GradExpn A B (Suc(Suc m))"
      using assms(2) assms(3) gradexp_stab_aux_n by blast 
    moreover
    have "(Suc (Suc m))  0"
      by simp 
    ultimately
    show ?thesis
      using GradExp_def by blast 
  qed
qed

lemma gradexp__grad_aux_1: 
  shows " C. (C = (GradExpn A A (Suc n))  Grad A A C)"
  by (simp add: grad_equiv_coq_1)

lemma gradexp__grad_aux: 
  assumes "A  B"
  shows " C. (C = (GradExpn A B (Suc n))  Grad A B C)" 
proof (induction n)
  case 0
  {
    fix C
    assume "C = (GradExpn A B (Suc 0))"
    hence "C = B"
      by force 
    hence "Grad A B C"
      using grad_equiv_coq_1 by auto 
  }
  hence "( C. (C = (GradExpn A B (Suc 0)))  Grad A B C)" 
    by blast
  thus ?case
    by blast 
next
  case (Suc n)
  {
    assume 1: " C. (C = (GradExpn A B (Suc n))  Grad A B C)"
    {
      fix C'
      assume "C' = (GradExpn A B ((Suc n) + 1))"
      have "A  B  ((Suc n) + 1)  0  ((Suc n)+1)  1"
        using assms by auto 
      then
      have "C' = (SymR A (GradExpn A B (Suc n)))"
        by (simp add: C' = GradExpn A B ((Suc n)+1)) 
      hence "Grad A B (GradExpn A B (Suc n))"
        using "1" by blast 
      have "Grad A B C'" 
      proof -
        have "Grad A B (GradExpn A B (Suc n))"
          using Grad A B (GradExpn A B (Suc n)) by blast 
        moreover
        have "Bet A (GradExpn A B (Suc n)) (GradExpn A B (Suc(Suc n)))" 
          using GradExpn_3 by blast 
        moreover
        have "Cong A (GradExpn A B (Suc n)) (GradExpn A B (Suc n)) C'"
          using GradExpn_4 C' = GradExpn A B (Suc n + 1) by force 
        ultimately
        show ?thesis
          using grad_sum by (metis Suc_eq_plus1 C' = GradExpn A B (Suc n + 1)) 
      qed
    }
    hence " C. (C = (GradExpn A B ((Suc n) + 1))  Grad A B C)"
      by blast
  }
  thus ?case
    using Suc.IH Suc_eq_plus1 by metis 
qed

lemma gradexp__grad: 
  assumes "GradExp A B C"
  shows "Grad A B C" 
proof -
  obtain n where "n  0  C = GradExpn A B n"
    using GradExp_def assms by blast
  hence "n  0" 
    by blast
  then obtain m where "n = Suc m"
    using not0_implies_Suc by presburger 
  thus ?thesis
    by (simp add: n  0  C = GradExpn A B n gradexp__grad_aux 
        gradexp__grad_aux_1)
qed

lemma gradexp_le__reach:
  assumes "GradExp A B B'" and
    "C D Le A B'"
  shows "Reach A B C D"
  using Reach_def assms(1) assms(2) gradexp__grad by blast 

lemma grad__ex_gradexp_le_aux_0:
  assumes "A = B"
  shows " D. GradExp A B D  A (Gradn A B n) Le A D" 
proof (induction n)
  case 0
  have "Gradn A B 0 = A"
    by auto
  moreover
  have "GradExp A B B"
    by (simp add: gradexp_init)
  moreover
  have "A A Le A B"
    by (simp add: le_trivial) 
  ultimately
  show " D. GradExp A B D  A (Gradn A B 0) Le A D"
    by auto 
next
  case (Suc n)
  {
    assume " D. GradExp A B D  A (Gradn A B n) Le A D" 
    then obtain D where "GradExp A B D  A (Gradn A B n) Le A D" 
      by blast
    let ?n = "Gradn A B n"
    let ?sucn = "Gradn A B (Suc n)"
    have "?n = ?sucn"
      using assms by force 
    hence " D. GradExp A B D  A (Gradn A B (Suc n)) Le A D"
      using D. GradExp A B D  A (Gradn A B n) Le A D by presburger 
  }
  thus ?case   
    using Suc.IH by blast 
qed

lemma grad__ex_gradexp_le_aux_1:
  assumes "A  B"
  shows " D. GradExp A B D  A (Gradn A B n) Le A D" 
proof (induction n)
  case 0
  have "Gradn A B 0 = A"
    by auto
  moreover
  have "GradExp A B B"
    by (simp add: gradexp_init)
  moreover
  have "A A Le A B"
    by (simp add: le_trivial) 
  ultimately
  show " D. GradExp A B D  A (Gradn A B 0) Le A D"
    by auto 
next
  case (Suc n)
  {
    assume " D. GradExp A B D  A (Gradn A B n) Le A D" 
    then obtain D where "GradExp A B D  A (Gradn A B n) Le A D" 
      by blast
    obtain D' where "Bet A D D'  Cong D D' A D"
      using segment_construction by blast
    have "(Gradn A B n) (Gradn A B (Suc n)) Le D D'"
      by (meson Cong_Gradn_Gradn_Suc Bet A D D'  Cong D D' A D 
          GradExp A B D  A (Gradn A B n) Le A D grad__le gradexp__grad 
          l5_6 le_right_comm not_cong_4312) 
    hence "GradExp A B D'  A (Gradn A B (Suc n)) Le A D'"
      by (meson Diff__Bet_Gradn_Gradn_Suc Bet A D D'  Cong D D' A D
          GradExp A B D  A (Gradn A B n) Le A D assms bet2_le2__le1346 
          gradexp_stab not_cong_3412) 
    hence " D. GradExp A B D  A (Gradn A B (Suc n)) Le A D" 
      by blast
  }
  thus ?case   
    using Suc.IH by blast 
qed

lemma grad__ex_gradexp_le_aux:
  shows " D. GradExp A B D  A (Gradn A B n) Le A D" 
  using grad__ex_gradexp_le_aux_0 grad__ex_gradexp_le_aux_1 by blast

lemma grad__ex_gradexp_le:
  assumes "Grad A B C"
  shows " D. GradExp A B D  A C Le A D" 
  using grad__ex_gradexp_le_aux Grad_def assms by auto 

lemma reach__ex_gradexp_le: 
  assumes "Reach A B C D"
  shows " B'. GradExp A B B'  C D Le A B'"
  by (meson Reach_def le_transitivity assms grad__ex_gradexp_le)

lemma gradexp2_init:
  shows "GradExp2 A B B C D D" 
proof -
  have "B = GradExpn A B (Suc 0)"
    by auto
  moreover
  have "D = GradExpn C D (Suc 0)"
    by auto
  ultimately
  show ?thesis
    using GradExp2_def by blast
qed

lemma GradExp2_stab:
  assumes "GradExp2 A B C D E F" and
    "Bet A C C'" and
    "Cong A C C C'" and
    "Bet D F F'" and
    "Cong D F F F'"
  shows "GradExp2 A B C' D E F'"
proof -
  obtain n where "(n  0)  (C = GradExpn A B n)  (F = GradExpn D E n)"
    using GradExp2_def assms(1) by presburger
  hence "n  0" 
    by blast
  then obtain m where "Suc m = n"
    using not0_implies_Suc by auto
  hence "C = GradExpn A B (Suc m)"
    using n  0  C = GradExpn A B n  F = GradExpn D E n by blast 
  hence "C' = GradExpn A B (Suc (Suc m))"
    using assms(2) assms(3) gradexp_stab_aux_n by blast 
  moreover
  have "F = GradExpn D E (Suc m)"
    using Suc m = n n  0  C = GradExpn A B n  F = GradExpn D E n by blast
  hence "F' = GradExpn D E (Suc (Suc m))"
    using assms(4) assms(5) gradexp_stab_aux_n by blast 
  moreover
  have "Suc (Suc m)  0" 
    by simp
  ultimately
  show ?thesis
    using GradExp2_def by blast 
qed

lemma gradexp2__gradexp123:
  assumes "GradExp2 A B C D E F"
  shows "GradExp A B C"
proof -
  obtain n where "n  0  (C = GradExpn A B n)  (F = GradExpn D E n)"
    using assms(1) GradExp2_def by blast
  thus ?thesis using GradExp_def by blast
qed

lemma gradexp2__gradexp456:
  assumes "GradExp2 A B C D E F"
  shows "GradExp D E F"
proof -
  obtain n where "n  0  (C = GradExpn A B n)  (F = GradExpn D E n)"
    using assms(1) GradExp2_def by blast
  thus ?thesis using GradExp_def by blast
qed

lemma MidR_uniq:
  assumes "C = MidR A B" and
    "D = MidR A B" 
  shows "C = D"
  by (simp add: assms(1) assms(2))

lemma MidR_ex_0:
  shows "(MidR A B) Midpoint A B" 
proof -
  have "MidR A B = (SOME x. x Midpoint A B)" 
    by simp
  thus ?thesis 
    using someI_ex by (metis midpoint_existence) 
qed

lemma MidR_ex_1:
  assumes "C = (MidR A B)" 
  shows "C Midpoint A B" 
  using assms MidR_ex_0 by blast

lemma MidR_ex_aux:
  assumes "C Midpoint A B"
  shows "C = (SOME x. x Midpoint A B)" 
  by (metis assms MidR_uniq_aux someI)   

lemma MidR_ex:
  assumes "C Midpoint A B"
  shows "C = (MidR A B)"
  using MidR_ex_0 MidR_uniq_aux assms by blast 

lemma gradexpinv_init_aux:
  shows "B = GradExpInvn A B 0"
  by simp

lemma gradexpinv_init:
  shows "GradExpInv A B B"
  using gradexpinv_init_aux using GradExpInv_def by blast 

lemma gradexpinv_stab_aux_0:
  assumes "B = GradExpInvn A C 0" and
    "Bet A B' B" and
    "Cong A B' B' B"
  shows "B' = GradExpInvn A C (Suc 0)" 
proof -
  have "B = C" 
    using assms(1) by simp
  have "B' Midpoint A B"
    by (simp add: assms(2) assms(3) midpoint_def) 
  hence "B' = MidR A B"
    using MidR_ex by blast
  show ?thesis
  proof (cases "A = C")
    case True
    hence "B' = B"
      using B = C assms(2) bet_neq12__neq by blast 
    moreover
    have "A = GradExpInvn A C (Suc 0)" 
      using True by simp
    ultimately
    show ?thesis
      using True B = C by force
  next
    case False
    hence "MidR A C = GradExpInvn A C (Suc 0)" 
      by simp
    thus ?thesis
      using B = C B' = MidR A B by blast 
  qed
qed

lemma gradexpinv_stab_aux_n:
  assumes "Bet A B' B" and
    "Cong A B' B' B" and
    "B = GradExpInvn A C (Suc n)"
  shows "B' = GradExpInvn A C (Suc (Suc n))" 
proof -
  have "Suc (Suc n)  0" by simp
  moreover
  have "Suc (Suc n)  1" by simp
  moreover
  have "B' Midpoint A B"
    using Midpoint_def assms(1) assms(2) by blast 
  hence "B' = MidR A (GradExpInvn A C (Suc n))"
    using MidR_ex assms(3) by blast 
  ultimately
  show ?thesis
    by (metis GradExpInvn.simps One_nat_def assms(1) assms(3) 
        diff_Suc_Suc l8_20_1_R1 l8_6 minus_nat.diff_0)  
qed

lemma gradexpinv_stab:
  assumes "Bet A B' B" and
    "Cong A B' B' B" and
    "GradExpInv A B C"
  shows "GradExpInv A B' C" 
proof -
  have "GradExpInv A B C"
    using GradExpInv_def assms(3) by auto
  then obtain n where "B = GradExpInvn A C n"
    using GradExpInv_def by blast
  have "B' = GradExpInvn A C (Suc n)" 
  proof (cases "n = 0")
    case True
    show ?thesis
      using True B = GradExpInvn A C n assms(1) assms(2) 
        gradexpinv_stab_aux_0 by blast 
  next
    case False
    then obtain k where "Suc k = n"
      using not0_implies_Suc by auto 
    hence "B = GradExpInvn A C (Suc k)" 
      using B = GradExpInvn A C n by blast
    hence "B' = GradExpInvn A C (Suc (Suc k))" 
      using gradexpinv_stab_aux_n assms(1) assms(2) by blast
    show ?thesis
      using B' = GradExpInvn A C (Suc (Suc k)) Suc k = n by blast 
  qed
  moreover
  have "Suc n  0"
    by simp 
  ultimately
  show ?thesis 
    using GradExpInv_def by blast
qed

lemma gradexp__gradexpinv_aux_1_0:
  assumes "C = GradExpn A B (Suc 0)"
  shows "B = GradExpInvn A C 0" 
proof (cases "A = B")
  case True
  hence "C = A"
    by (simp add: assms)
  thus ?thesis
    by (simp add: True)
next
  case False
  hence "C = B" 
    using assms by simp
  thus ?thesis 
    using False by simp
qed

lemma SymR_MidR:
  assumes "A = SymR B C"
  shows "C = MidR A B"
proof -
  have "C Midpoint B A" 
    using assms SymR_ex symmetric_point_construction by blast 
  thus ?thesis
    using MidR_ex_0 l7_17_bis by auto 
qed

lemma MidR_comm:
  assumes "C = MidR A B"
  shows "C = MidR B A"
proof -
  have "C Midpoint B A"
    using MidR_ex_0 assms l7_2 by blast
  thus ?thesis
    by (meson MidR_ex C Midpoint B A) 
qed

lemma MidR_SymR:
  assumes "C = MidR A B"
  shows "A = SymR B C" 
proof -
  have "C Midpoint B A"
    using MidR_ex_0 assms l7_2 by blast
  thus ?thesis
    using SymR_ex by auto 
qed

lemma MidR_AA:
  shows "A = MidR A A" 
  using l7_3_2 MidR_ex by blast 

lemma MidR_AB:
  assumes "A = MidR A B"
  shows "A = B"
  by (metis MidR_AA MidR_SymR MidR_comm assms) 

lemma gradexp__gradexpinv_aux_1_n_aa:
  shows "GradExpInvn A' (MidR A' C') n = GradExpInvn A' C' (Suc n)" 
proof (cases "A' = C'")
  case True
  hence "A' = MidR A' C'" 
    using l7_3_2 MidR_ex by blast 
  hence "GradExpInvn A' (MidR A' C') n = A'"
    by simp
  moreover
  have "GradExpInvn A' C' (Suc n) = A'"
    using True by simp
  ultimately
  show ?thesis 
    by simp
next
  case False
  have "GradExpInvn A' (MidR A' C') n = GradExpInvn A' C' (Suc n)" 
  proof (induction n)
    case 0
    hence "A'  MidR A' C'"
      using False MidR_AB by blast 
    hence "GradExpInvn A' (MidR A' C') 0 = (MidR A' C')"
      by simp 
    moreover
    have "GradExpInvn A' C' 0 = C'"
      using gradexpinv_init_aux by presburger 
    hence "GradExpInvn A' C' (Suc 0) = (MidR A' C')" 
      using False by simp
    ultimately
    have "GradExpInvn A' (MidR A' C') 0 = GradExpInvn A' C' (Suc 0)" 
      by simp
    thus ?case 
      by blast
  next
    case (Suc n)
    {
      assume "GradExpInvn A' (MidR A' C') n = GradExpInvn A' C' (Suc n)" 
      have "GradExpInvn A' (MidR A' C') (Suc n) 
    = (MidR A' (GradExpInvn A' (MidR A' C') ((Suc n)-1)))"
        by (metis GradExpInvn.simps MidR_AB Suc_neq_Zero diff_self_eq_0)
      have " = MidR A' (GradExpInvn A' (MidR A' C') n)" 
        by simp
      have " = MidR A' (GradExpInvn A' C' (Suc n))"
        using Suc.IH by presburger 
      have " = MidR A' (GradExpInvn A' C' ((Suc(Suc n))-1))"
        using diff_Suc_1 by presburger 
      have " = GradExpInvn A' C' (Suc(Suc n))"
        using False by auto 
      hence "GradExpInvn A' (MidR A' C') (Suc n) = GradExpInvn A' C' (Suc(Suc n))"
        using GradExpInvn A' (MidR A' C') (Suc n) 
      = MidR A' (GradExpInvn A' (MidR A' C') (Suc n - 1)) 
          MidR A' (GradExpInvn A' (MidR A' C') (Suc n - 1)) 
      = MidR A' (GradExpInvn A' (MidR A' C') n) 
          MidR A' (GradExpInvn A' (MidR A' C') n) 
      = MidR A' (GradExpInvn A' C' (Suc n)) 
          MidR A' (GradExpInvn A' C' (Suc n)) 
      = MidR A' (GradExpInvn A' C' (Suc (Suc n) - 1)) 
        by presburger      
    }
    thus ?case
      using Suc.IH by blast 
  qed
  thus ?thesis by blast
qed

lemma gradexp__gradexpinv_aux_1_n_a:
  assumes " A B C. (C = GradExpn A B (Suc n))  B = GradExpInvn A C n"
  shows " A' B' C'. (C' = GradExpn A' B' (Suc(Suc n)))  B' = GradExpInvn A' C' (Suc n)"
proof -
  {
    fix A' B' C'
    assume "C' = GradExpn A' B' (Suc(Suc n))"
    have "B' = GradExpInvn A' C' (Suc n)"
    proof (cases "A' = B'")
      case True
      {
        hence "C' = A'" 
          using True by (simp add: C' = GradExpn A' B' (Suc (Suc n))) 
        hence "A' = GradExpInvn A' C' (Suc n)"
          by simp
        hence "B' = GradExpInvn A' C' (Suc n)" 
          using True by simp
      }
      thus ?thesis
        using C' = GradExpn A' B' (Suc (Suc n)) by blast 
    next
      case False
      {
        have "Suc (Suc n)  0" 
          by simp
        moreover
        have "Suc (Suc n)  1" 
          by simp
        moreover
        have "C' = SymR A' (GradExpn A' B' ((Suc (Suc n))-1))"
          by (simp add: False C' = GradExpn A' B' (Suc (Suc n)))
        have "(Suc (Suc n))-1 = Suc n" 
          by simp
        let ?ssn1 = "(Suc (Suc n))-1"
        let ?B1 = "GradExpn A' B' ?ssn1"
        have "?B1 = (GradExpn A' B' (Suc n))"
          by simp 
        hence "B' = (GradExpInvn A' ?B1 n)" 
          using assms by blast
        have "C' = SymR A' ?B1"
          using C' = SymR A' (GradExpn A' B' (Suc (Suc n) - 1)) by blast 
        {
          assume "n = 0"
          hence "B' = ?B1"
            by force 
        }
        {
          assume "n = 1"
          hence "B' = MidR A' ?B1"
            by (metis False GradExpInvn.simps One_nat_def 
                B' = GradExpInvn A' (GradExpn A' B' (Suc (Suc n) - 1)) n 
                GradExpn A' B' (Suc (Suc n) - 1) = GradExpn A' B' (Suc n) 
                n_not_Suc_n) 
        }
        {
          assume "n  0  n  1"
          hence "B' = MidR A' (GradExpInvn A' ?B1 (n-1))"
            by (metis False GradExpInvn.simps 
                B' = GradExpInvn A' (GradExpn A' B' (Suc (Suc n) - 1)) n) 
        }
        {
          assume "Suc n = 1"
          hence "n = 0" 
            by simp
          have "C' = SymR A' ?B1"
            using C' = SymR A' (GradExpn A' B' (Suc (Suc n) - 1)) by blast 
          hence "?B1 = MidR C' A'"
            using SymR_MidR by blast 
          hence "?B1 = MidR A' C'"
            using MidR_comm by blast 
          hence "B' = MidR A' C'"
            using n = 0  B' = GradExpn A' B' (Suc (Suc n) - 1) n = 0 
            by force 
        }
        moreover
        {
          assume "Suc n  1"
          hence "B' = MidR A' (GradExpInvn A' ?B1 (n-1))"
            by (metis One_nat_def 
                n = 1  B' = MidR A' (GradExpn A' B' (Suc (Suc n) - 1)) 
                n  0  n  1  B' 
            = MidR A' (GradExpInvn A' (GradExpn A' B' (Suc (Suc n) - 1)) (n - 1)) 
                diff_Suc_1 gradexpinv_init_aux) 
          obtain k where "Suc k = n"
            using Suc n  1 not0_implies_Suc by auto 
          have "C' = SymR A' ?B1"
            using C' = SymR A' (GradExpn A' B' (Suc (Suc n) - 1)) by blast 
          hence "?B1 = MidR A' C'"
            using MidR_comm SymR_MidR by presburger 
          hence "GradExpInvn A' ?B1 k = (GradExpInvn A' C' (Suc k))" 
            using gradexp__gradexpinv_aux_1_n_aa by presburger 
          hence "GradExpInvn A' ?B1 (n-1) = (GradExpInvn A' C' n)"
            using Suc k = n by force 
          hence "B' = MidR A' (GradExpInvn A' C' n)"
            using B' = MidR A' (GradExpInvn A' (GradExpn A' B' (Suc (Suc n) - 1)) (n - 1)) 
            by presburger 
          hence "B' = MidR A' (GradExpInvn A' C' ((Suc n)-1))"
            using diff_Suc_1 by presburger 
        }
        ultimately
        have "B' = GradExpInvn A' C' (Suc n)"
          by (metis GradExpInvn.elims GradExpn_2 Suc_neq_Zero 
              C' = GradExpn A' B' (Suc (Suc n)) bet_neq12__neq) 
      }
      hence "(C' = GradExpn A' B' (Suc(Suc n)))  B' = GradExpInvn A' C' (Suc n)" by blast
      thus ?thesis
        using C' = GradExpn A' B' (Suc (Suc n)) by blast 
    qed
  }
  thus ?thesis by blast
qed

lemma gradexp__gradexpinv_aux_1_b:
  shows " A B C. C = GradExpn A B (Suc n)  B = GradExpInvn A C n" 
proof (induction n)
  case 0
  have " A B C. C = GradExpn A B (Suc 0)  B = GradExpInvn A C 0"
    by (meson gradexp__gradexpinv_aux_1_0) 
  thus ?case by blast
next
  case (Suc n)
  {
    assume  " A B C. C = GradExpn A B (Suc n)  B = GradExpInvn A C n" 
    {
      fix A B C
      assume "C = GradExpn A B (Suc(Suc n))" 
      have "B = GradExpInvn A C (Suc n)"
        using C = GradExpn A B (Suc (Suc n)) 
          A B C. C = GradExpn A B (Suc n)  B = GradExpInvn A C n 
          gradexp__gradexpinv_aux_1_n_a by presburger 
    }
    hence " A B C. C = GradExpn A B (Suc(Suc n))  B = GradExpInvn A C (Suc n)" 
      by blast
  }
  thus ?case
    using Suc.IH by fastforce 
qed

lemma gradexp__gradexpinv_aux_1:
  assumes "C = GradExpn A B (Suc n)"
  shows "B = GradExpInvn A C n"
  using assms gradexp__gradexpinv_aux_1_b by blast 

lemma gradexp__gradexpinv_aux:
  assumes "GradExp A B C"
  shows "GradExpInv A B C" 
proof -
  obtain n where "n  0  C = GradExpn A B n" 
    using assms GradExp_def by blast
  hence "n  0"
    by blast
  then obtain m where "n = Suc m"
    using not0_implies_Suc by presburger 
  hence "B = GradExpInvn A C m" 
    using gradexp__gradexpinv_aux_1
    using n  0  C = GradExpn A B n by blast 
  thus ?thesis
    using GradExpInv_def n  0  C = GradExpn A B n by blast 
qed

lemma gradexpinv__gradexp_aux_1_a_0: 
  assumes "B' = MidR A' (GradExpInvn A' C' (Suc 0))"
  shows "C' = SymR A' (GradExpn A' B' (Suc(Suc 0)))" 
proof (cases "A' = C'")
  case True
  hence "B' = MidR A' C'"
    by (simp add: assms) 
  thus ?thesis
    by (metis GradExpn_1 MidR_AA MidR_SymR True)
next
  case False
  thus ?thesis
    by (metis GradExpn_2 GradExpn_4 Mid_cases SymR_MidR 
        assms gradexp__gradexpinv_aux_1_0 gradexp__gradexpinv_aux_1_n_aa 
        gradexpinv_init_aux gradexpinv_stab_aux_0 l7_9 
        midpoint_def midpoint_existence) 
qed

lemma sym_mid:
  shows "SymR A (MidR A B) = B"
  by (metis MidR_SymR MidR_comm)

lemma gradexpn_suc_suc:
  shows "GradExpn A B (Suc n) = GradExpn A (MidR A B) (Suc(Suc n))" 
proof (cases "A = B")
  case True 
  thus ?thesis
    using GradExpn_1 MidR_AA by presburger 
next
  case False
  hence "A  B" 
    by simp
  have  "GradExpn A B (Suc n) = GradExpn A (MidR A B) (Suc(Suc n))" 
  proof (induction n)
    case 0
    have "GradExpn A (MidR A B) (Suc(Suc 0)) 
  = SymR A (GradExpn A (MidR A B) ((Suc(Suc 0))-1))"
      using False MidR_AB 
      by (metis diff_Suc_1 gradexp__gradexpinv_aux_1_b 
          gradexp__gradexpinv_aux_1_n_aa gradexpinv_init_aux sym_mid)
    hence " = SymR A (GradExpn A (MidR A B) (Suc 0))"
      using diff_Suc_1 by presburger
    hence " = SymR A (MidR A B)"
      by simp
    hence " = B"
      using sym_mid by blast
    hence "GradExpn A B (Suc 0) = GradExpn A (MidR A B) (Suc (Suc 0))"
      by (metis gradexp__gradexpinv_aux_1 gradexp__gradexpinv_aux_1_n_aa 
          gradexpinv__gradexp_aux_1_a_0 gradexpinv_init_aux) 
    thus ?case
      by blast 
  next
    case (Suc n)
    {
      assume " A' B'. GradExpn A' B' (Suc n) = GradExpn A' (MidR A' B') (Suc(Suc n))"
      {
        fix A B
        have "GradExpn A (MidR A B) (Suc(Suc(Suc n))) 
                  = SymR A (GradExpn A (MidR A B) (Suc(Suc(Suc n))-1))"
          using MidR_AA MidR_SymR by fastforce
        hence " = GradExpn A B (Suc(Suc n))"
          by (metis GradExpn_3 GradExpn_4 
              A' B'. GradExpn A' B' (Suc n) = GradExpn A' (MidR A' B') (Suc (Suc n)) 
              gradexp_stab_aux_n) 
        hence "GradExpn A B (Suc(Suc n)) = GradExpn A (MidR A B) (Suc(Suc(Suc n)))"
          using GradExpn A (MidR A B) (Suc (Suc (Suc n))) 
                    = SymR A (GradExpn A (MidR A B) (Suc (Suc (Suc n)) - 1)) 
          by presburger 
      }
      hence " A B. GradExpn A B (Suc(Suc n)) = GradExpn A (MidR A B) (Suc(Suc(Suc n)))"
        by blast
    }
    thus ?case
      using GradExpn_3 GradExpn_4 Suc.IH gradexp_stab_aux_n by blast 
  qed
  thus ?thesis by blast
qed

lemma gradexpinv__gradexp_aux_1_a_n: 
  assumes " A B C. (B = MidR A (GradExpInvn A C (Suc n)))  
  C = SymR A (GradExpn A B (Suc(Suc n)))" 
  shows "B' = MidR A' (GradExpInvn A' C' (Suc(Suc n)))  
  C' = SymR A' (GradExpn A' B' (Suc(Suc(Suc n))))" 
proof (cases "A' = C'")
  case True
  thus ?thesis
    by (metis GradExpn_1 MidR_AA MidR_SymR gradexp__gradexpinv_aux_1) 
next
  case False
  {
    assume "B' = MidR A' (GradExpInvn A' C' (Suc(Suc n)))"
    hence " =  MidR A' (MidR A' (GradExpInvn A' C' (Suc(Suc n)-1)))"
      using False by force
    hence " = MidR A' (MidR A' (GradExpInvn A' C' (Suc n)))"
      by simp
    let ?B1 = "(MidR A' (GradExpInvn A' C' (Suc n)))"
    have "C' = SymR A' (GradExpn A' ?B1 (Suc(Suc n)))" 
      using assms by blast
    have "GradExpn A' ?B1 (Suc(Suc n)) = GradExpn A' (MidR A' ?B1) (Suc(Suc(Suc n)))" 
      using gradexpn_suc_suc by blast 
    hence "C' = SymR A' (GradExpn A' B' (Suc(Suc(Suc n))))"
      by (metis MidR_comm SymR_MidR B' = MidR A' (GradExpInvn A' C' (Suc (Suc n))) 
          C' = SymR A' (GradExpn A' (MidR A' (GradExpInvn A' C' (Suc n))) (Suc (Suc n))) 
          gradexp__gradexpinv_aux_1 gradexp__gradexpinv_aux_1_n_aa)
  }
  thus ?thesis 
    by blast
qed

lemma gradexpinv__gradexp_aux_1_a:
  shows " A B C. B = MidR A (GradExpInvn A C (Suc n))  
  C = SymR A (GradExpn A B (Suc(Suc n)))" 
proof (induction n)
  case 0
  thus ?case
    using gradexpinv__gradexp_aux_1_a_0 by blast
next
  case (Suc n)
  thus ?case
    using gradexpinv__gradexp_aux_1_a_n by blast
qed

lemma gradexpinv__gradexp_aux_1_n:
  assumes "B = GradExpInvn A C n  C = GradExpn A B (Suc n)" 
  shows "B' = GradExpInvn A' C' (Suc n)C' = GradExpn A' B' (Suc(Suc n))"
proof (cases "A' = C'")
  case True
  thus ?thesis 
    by force
next
  case False
  hence "A'  C'" 
    by blast
  {
    assume "B' = GradExpInvn A' C' (Suc n)" 
    have "C' = GradExpn A' B' (Suc(Suc n))"
    proof (cases "Suc n = 1")
      case True
      hence "B' = GradExpInvn A' C' 1" 
        using B' = GradExpInvn A' C' (Suc n) by simp 
      hence "B' = MidR A' C'" 
        using A'  C' by simp
      hence "C' = SymR A' B'"
        by (metis MidR_SymR MidR_comm)
      have "A'  B'"
        using False MidR_AB B' = MidR A' C' by blast 
      have "Suc(Suc n)  0  Suc (Suc n)  1" 
        by simp
      hence "GradExpn A' B' (Suc(Suc n)) =  SymR A' (GradExpn A' B' (Suc(Suc n)-1))"
        by (simp add: A'  B') 
      hence " = SymR A' B'"
        by (simp add: True) 
      thus ?thesis
        using C' = SymR A' B' 
          GradExpn A' B' (Suc (Suc n)) = SymR A' (GradExpn A' B' (Suc (Suc n) - 1)) 
        by presburger 
    next
      case False
      hence "B' = (MidR A' (GradExpInvn A' C' ((Suc n)-1)))"
        using B' = GradExpInvn A' C' (Suc n) 
        by (simp add: A'  C') 
      hence "C' = SymR A' (GradExpn A' B' ((Suc(Suc n)-1)))" 
        using gradexpinv__gradexp_aux_1_a 
        by (metis GradExpn.simps One_nat_def A'  C' 
            B' = GradExpInvn A' C' (Suc n) gradexp__gradexpinv_aux_1 
            nat.distinct(1) nat.inject) 
      hence " = GradExpn A' B' (Suc(Suc n))"
        by (metis GradExpn.simps GradExpn_1 MidR_AB MidR_SymR 
            One_nat_def B' = MidR A' (GradExpInvn A' C' (Suc n - 1)) 
            nat.distinct(1) nat.inject) 
      thus ?thesis
        using C' = SymR A' (GradExpn A' B' (Suc (Suc n) - 1)) by blast 
    qed
  }
  thus ?thesis 
    by blast
qed

lemma gradexpinv__gradexp_aux_1:
  shows "B = GradExpInvn A C n  C = GradExpn A B (Suc n)" 
proof (induction n)
  case 0
  thus ?case by force
next
  case (Suc n)
  thus ?case 
    using gradexpinv__gradexp_aux_1_n by blast
qed

lemma gradexpinv__gradexp_aux:
  assumes "GradExpInv A B C"
  shows "GradExp A B C" 
proof -
  obtain n where "B = GradExpInvn A C n"
    using GradExpInv_def assms by blast
  hence "C = GradExpn A B (Suc n)" 
    using gradexpinv__gradexp_aux_1 by blast
  thus ?thesis
    using GradExp_def by blast 
qed

lemma gradexp__gradexpinv: 
  shows "GradExp A B C  GradExpInv A B C"
  using gradexp__gradexpinv_aux gradexpinv__gradexp_aux by blast 


lemma reach__ex_gradexp_lt_aux:
  shows " A B C P Q R. ((A  B  A B Le P R  R = GradExpn P Q (Suc n))  
  ( C. GradExp A C B  A C Lt P Q))" 
proof (induction n)
  case 0
  {
    fix A B C P Q R
    assume 1: "A  B" and 2:"A B Le P R" and 3: "R = GradExpn P Q (Suc 0)"
    obtain C where "Bet A C B  Cong A C C B"
      by (meson midpoint_bet midpoint_cong midpoint_existence)
    have "P  Q"
      using "1" "2" "3" le_zero by force 
    have "R = Q"
      by (simp add: "3") 
    have "GradExpInv A C B" 
    proof -
      have "GradExpInv A B B"
        using gradexpinv_init by blast
      thus ?thesis
        using Bet A C B  Cong A C C B gradexpinv_stab by blast 
    qed
    hence "GradExp A C B"
      by (simp add: gradexpinv__gradexp_aux) 
    moreover
    have "A C Lt P Q" 
    proof -
      have "A C Lt A B"
        by (simp add: "1" Bet A C B  Cong A C C B mid__lt midpoint_def) 
      moreover
      have "A B Le P Q"
        using "2" R = Q by auto
      ultimately
      show ?thesis
        using le3456_lt__lt by blast
    qed
    ultimately
    have " C. GradExp A C B  A C Lt P Q"
      by blast 
  }
  thus ?case by blast
next
  case (Suc n)
  {
    assume H1: " A' B' C' P' Q' R'. 
    (A'  B'  A' B' Le P' R'  R' = GradExpn P' Q' (Suc n) 
     ( C'. GradExp A' C' B'  A' C' Lt P' Q'))" 
    {
      fix A B C P Q R
      assume 1: "A  B" and 
        2:"A B Le P R" and 
        3: "R = GradExpn P Q (Suc (Suc n))"
      obtain M where "M Midpoint A B"
        using midpoint_existence by blast
      have "P  R"
        using "1" "2" le_diff by blast 
      have "M  A"
        using "1" M Midpoint A B bet_cong_eq between_trivial2 
          midpoint_cong by blast 
      have "M  B"
        using "1" M Midpoint A B cong_identity midpoint_cong by blast 
      have "A M Le P R"
        using "1" "2" M Midpoint A B le3456_lt__lt lt__le mid__lt by blast 
      have "R = SymR P (GradExpn P Q ((Suc(Suc n))-1))"
        using "3" P  R by auto
      hence " = SymR P (GradExpn P Q (Suc n))"
        using diff_Suc_1 by presburger 
      let ?R' = "GradExpn P Q (Suc n)"
      have " C. (GradExp A C M  A C Lt P Q)"
      proof -
        have "A M Le P ?R'"
          using "2" "3" GradExpn_3 GradExpn_4 M Midpoint A B 
            le_mid2__le12 midpoint_def by blast 
        moreover
        have "M  A" 
          using M  A by force
        moreover
        have "?R' = GradExpn P Q (Suc n)" 
          by simp
        ultimately
        show ?thesis 
          using H1 by force 
      qed
      then obtain C where "GradExp A C M  A C Lt P Q" 
        by auto
      hence " C. GradExp A C B  A C Lt P Q"
        using M Midpoint A B gradexp_stab midpoint_bet 
          midpoint_cong by blast 
    }
    hence " A B C P Q R. 
    A  B  A B Le P R  R = GradExpn P Q (Suc (Suc n)) 
     ( C. GradExp A C B  A C Lt P Q)" 
      by blast
  }
  thus ?case
    using Suc.IH by presburger
qed

lemma reach__grad_min_1:
  assumes "A  B" and
    "Bet A B C"  and
    "A C Le A (Gradn A B (Suc 0))" 
  shows " D E. (Bet A D C  Grad A B D  E  C  Bet A C E  Bet A D E  Cong A B D E)" 
proof (cases "C = (Gradn A B (Suc 0))")
  case True
  let ?f = "Gradn A B (Suc (Suc 0))"
  have "Grad A B B"
    by (simp add: grad_equiv_coq_1) 
  moreover
  have "?f  C"
    by (metis Gradn_uniq_aux_1 True assms(1)) 
  moreover
  have "Bet A C ?f"
    using Bet_Gradn_Gradn_Suc True by blast 
  moreover
  have "Bet A B ?f"
    by (meson Diff__Bet_Gradn_Suc assms(1)) 
  moreover
  have "Cong A B B ?f"
    by (metis Lem_Gradn_0 Cong_Gradn_Gradn_Suc True assms(2) between_cong) 
  ultimately
  show ?thesis
    using assms(2) by blast
next
  case False
  let ?e = "Gradn A B (Suc 0)"
  have "Grad A B B"
    by (simp add: grad_equiv_coq_1) 
  moreover
  have "?e  C"
    using False by auto 
  moreover
  have "Bet A C ?e"
    by (meson Gradn_aux_1_0 Diff__Bet_Gradn_Suc assms(1) 
        assms(2) assms(3) bet_out l5_1 l6_13_1 l6_6) 
  moreover
  have "Bet A B ?e"
    using Diff__Bet_Gradn_Suc not_bet_distincts by blast 
  moreover
  have "Cong A B B ?e"
    by (metis (full_types) False Lem_Gradn_0 Cong_Gradn_Gradn_Suc 
        assms(2) bet_neq23__neq between_cong_2 between_exchange3 
        calculation(3) calculation(4) l5_1)
  ultimately
  show ?thesis
    using assms(2) by blast
qed

lemma reach__grad_min_n:
  assumes "A  B" and 
    "Bet A B C" and 
    "A C Le A (Gradn A B (Suc n))  
  ( D E. (Bet A D C  Grad A B D  E  C  Bet A C E  Bet A D E  Cong A B D E))"
  shows "A C Le A (Gradn A B (Suc(Suc n)))  
  ( D E. (Bet A D C  Grad A B D  E  C  Bet A C E  Bet A D E  Cong A B D E))"
proof -
  {
    assume
      H1: "A C Le A (Gradn A B (Suc(Suc n)))"
    have  " D E. (Bet A D C  Grad A B D  E  C  Bet A C E  Bet A D E  Cong A B D E)" 
    proof (cases "A C Le A (Gradn A B (Suc n))")
      case True
      hence "A C Le A (Gradn A B (Suc n))"
        by blast
      then obtain D E where 
        "Bet A D C  Grad A B D  E  C  Bet A C E  Bet A D E  Cong A B D E" 
        using assms(1) assms(3) by blast 
      thus ?thesis 
        by blast
    next
      case False
      hence "A (Gradn A B (Suc n)) Lt A C"
        by (meson nlt__le)
      let ?f = "Gradn A B (Suc (Suc n))"
      show ?thesis 
      proof (cases "C = ?f")
        case True
        let ?e = "Gradn A B (Suc (Suc (Suc n)))"
        show ?thesis 
        proof -
          have "Bet A C C"
            using not_bet_distincts by auto 
          moreover
          have "C = Gradn A B (Suc (Suc n))"
            using True by auto 
          hence "Grad A B C"
            using Grad_def by blast 
          moreover
          have "?e  C"
            by (metis Gradn_uniq_aux_1 True assms(1)) 
          moreover
          have "Bet A C ?e"
            using Bet_Gradn_Gradn_Suc True by blast 
          moreover
          have "Bet A B ?e"
            using Diff__Bet_Gradn_Suc not_bet_distincts by blast 
          moreover
          have "Cong A B C ?e"
            using Cong_Gradn_Gradn_Suc True by blast
          ultimately
          show ?thesis
            using assms(2) by blast
        qed
      next
        case False 
        hence " A C Le A ?f"
          using H1 by blast
        show ?thesis
        proof -
          let ?d = "Gradn A B (Suc n)"
          let ?e = "Gradn A B (Suc(Suc n))"
          have "Bet A ?d C"
            by (meson Diff__Bet_Gradn_Suc 
                A (Gradn A B (Suc n)) Lt A C assms(1) assms(2)
                l5_1 l5_12_a lt__nle) 
          moreover
          have "Grad A B ?d"
            using Grad_def by blast 
          moreover
          have "?e  C"
            using False by auto
          moreover
          have "Bet A C ?e"
            by (metis H1 Bet_Gradn_Gradn_Suc assms(1) bet__lt1213 
                calculation(1) calculation(2) grad_neq__neq13 l5_1 lt__nle)
          moreover
          have "Bet A B ?e"
            using Diff__Bet_Gradn_Suc not_bet_distincts by blast 
          moreover
          have "Cong A B ?d ?e"
            using Cong_Gradn_Gradn_Suc by blast
          ultimately
          show ?thesis
            using between_exchange4 by blast
        qed
      qed
    qed
  }
  thus ?thesis
    by blast 
qed

lemma reach__grad_min_aux:
  assumes "A  B" and
    "Bet A B C"
  shows "(Grad A B (Gradn A B (Suc n))  A C Le A (Gradn A B (Suc n))) 
   ( D E. (Bet A D C  Grad A B D  E  C  Bet A C E  Bet A D E  Cong A B D E))"
proof (induction n)
  case 0
  show ?case   
    using assms(1) assms(2) reach__grad_min_1 by blast
next
  case (Suc n)
  have "Grad A B (Gradn A B (Suc n))  Grad A B (Gradn A B (Suc (Suc n)))"
    using Grad_def by blast
  thus ?case 
    using assms(1) assms(2) reach__grad_min_n Suc.IH by blast
qed

(** D is the last graduation of AB before or equal to C, and E the first graduation after C *)
lemma reach__grad_min:
  assumes "A  B" and
    "Bet A B C" and
    "Reach A B A C"
  shows " D E. (Bet A D C  Grad A B D  E  C  Bet A C E  Bet A D E  Cong A B D E)" 
proof -
  obtain D where "Grad A B D  A C Le A D"
    by (meson assms(3) gradexp__grad reach__ex_gradexp_le)
  hence "A C Le A D" 
    by blast
  have "A Out C D" 
  proof -
    have "A Out C B"
      using assms(1) assms(2) bet_out l6_6 by presburger 
    moreover
    have "Grad A B D"
      by (simp add: Grad A B D  A C Le A D)
    hence "Bet A B D"
      by (simp add: grad__bet) 
    hence "A Out B D"
      using assms(1) bet_out by force 
    ultimately
    show ?thesis
      using l6_7 by blast
  qed
  hence "Bet A C D" 
    using l6_13_1 by (simp add: A C Le A D) 
  have "Grad A B D"
    by (simp add: Grad A B D  A C Le A D)
  then obtain n where "(n  0)  (D = Gradn A B n)"
    using Grad_def by blast
  hence "n  0" 
    by simp
  then obtain m where "Suc m = n"
    by (metis not0_implies_Suc) 
  hence "D = Gradn A B (Suc m)"
    by (simp add: n  0  D = Gradn A B n) 
  show ?thesis 
    by (metis Grad A B D  A C Le A D 
        Suc m = n n  0  D = Gradn A B n 
        assms(1) assms(2) reach__grad_min_aux) 
qed

lemma reach__ex_gradexp_lt:
  assumes "A  B" and
    "Reach P Q A B"
  shows " C. GradExp A C B  A C Lt P Q"
proof -
  obtain R where "GradExp P Q R  A B Le P R"
    using assms(2) reach__ex_gradexp_le by blast
  then obtain n where "n  0  R = GradExpn P Q n"
    using GradExp_def by blast 
  hence "n  0" 
    by blast
  then obtain m where "Suc m = n"
    using not0_implies_Suc by blast
  hence "R = GradExpn P Q (Suc m)"
    by (simp add: n  0  R = GradExpn P Q n) 
  {
    fix B 
    assume "A  B" and "A B Le P R" 
    hence " C. (GradExp A C B  A C Lt P Q)"
      using R = GradExpn P Q (Suc m) reach__ex_gradexp_lt_aux by force
  }
  thus ?thesis
    by (simp add: GradExp P Q R  A B Le P R assms(1)) 
qed

(** This development is inspired by The Foundations of Geometry and 
  the Non-Euclidean Plane, by George E Martin, chapter 22 *)

lemma t22_18_aux_0:
  assumes "Bet A0 D1 A1" and
    "Cong E0 E1 A1 D1" and
    "D = Gradn A0 D1 (Suc 0)"
  shows " A E. (Grad2 A0 A1 A E0 E1 E  Cong E0 E A D  Bet A0 D A)" 
proof -
  have "D = D1" 
    by (simp add: assms(3))  
  thus ?thesis 
    using assms(1) assms(2) grad2_init by blast
qed

lemma t22_18_aux_n:
  assumes " A0 D1 A1 E0 E1 D. 
  (Bet A0 D1 A1  Cong E0 E1 A1 D1  D = Gradn A0 D1 (Suc n)) 
  ( A E. (Grad2 A0 A1 A E0 E1 E  Cong E0 E A D  Bet A0 D A))" 
  shows " A0 D1 A1 E0 E1 D. 
  (Bet A0 D1 A1  Cong E0 E1 A1 D1  D = Gradn A0 D1 (Suc(Suc n))) 
  ( A E. (Grad2 A0 A1 A E0 E1 E  Cong E0 E A D  Bet A0 D A))" 
proof -
  have " A0 D1 A1 E0 E1 D. (Bet A0 D1 A1  Cong E0 E1 A1 D1 
  D = Gradn A0 D1 (Suc n)) 
  ( A E. (Grad2 A0 A1 A E0 E1 E  Cong E0 E A D  Bet A0 D A))" 
    using assms by blast
  {
    fix A0 D1 A1 E0 E1 D
    assume 1: "Bet A0 D1 A1" and
      2: "Cong E0 E1 A1 D1" and
      "D = Gradn A0 D1 (Suc(Suc n))"
    let ?C = "Gradn A0 D1 (Suc n)"
    obtain A E where "Grad2 A0 A1 A E0 E1 E" and "Cong E0 E A ?C" and "Bet A0 ?C A" 
      using 1 2 assms by blast
    obtain A' where "Bet A0 A A'" and "Cong A A' A0 A1" 
      using segment_construction by blast
    obtain E' where "Bet E0 E E'" and "Cong E E' E0 E1"
      using segment_construction by blast
    have "Grad2 A0 A1 A' E0 E1 E'" 
      using Grad2_stab Bet A0 A A' Bet E0 E E' Cong A A' A0 A1 
        Cong E E' E0 E1 Grad2 A0 A1 A E0 E1 E cong_symmetry by blast
    moreover
    have "E0 E1 Le A A'" 
      by (meson "1" "2" Cong A A' A0 A1 bet__le2313 l5_6 
          le_left_comm not_cong_3412)
    obtain D' where "Bet A D' A'" and "Cong E0 E1 A D'"
      using Le_def E0 E1 Le A A' by blast
    have "Cong E0 E' A' D" 
    proof -
      have "Cong E0 E' ?C D'" 
      proof -
        have "Bet ?C A D'" 
          by (metis between_exchange3 Bet A D' A' 
              Bet A0 (Gradn A0 D1 (Suc n)) A Bet A0 A A' 
              between_inner_transitivity)
        moreover
        have "Cong E0 E ?C A" 
          using Cong E0 E A (Gradn A0 D1 (Suc n)) not_cong_1243 by blast
        moreover
        have "Cong E E' A D'" 
          by (metis cong_transitivity Cong E E' E0 E1 Cong E0 E1 A D')
        ultimately
        show ?thesis 
          using Bet E0 E E' l2_11 by blast
      qed
      moreover
      have "Cong ?C D' A' D" 
      proof -
        have "Bet ?C D' A'" 
          by (metis between_exchange3 Bet A D' A' 
              Bet A0 (Gradn A0 D1 (Suc n)) A Bet A0 A A' between_exchange2)
        moreover
        have "Bet A' D ?C" 
        proof -
          have "Grad A0 A1 A" 
            using Grad2 A0 A1 A E0 E1 E grad2__grad123 by blast
          hence "Bet A0 A1 A" 
            by (simp add: grad__bet)
          have "Bet A0 D1 ?C" 
            using Diff__Bet_Gradn_Suc not_bet_distincts by blast
          show ?thesis 
          proof (cases "A0 = ?C")
            case True
            thus ?thesis 
              by (metis Gradn_aux_1_0 Lem_Gradn_id_n 
                  D = Gradn A0 D1 (Suc (Suc n)) not_bet_distincts)
          next
            case False
            hence "A0  ?C" 
              by blast
            show ?thesis 
            proof (cases "D = ?C")
              case True
              thus ?thesis 
                using not_bet_distincts by blast
            next
              case False
              hence "D  ?C"
                by blast
              show ?thesis 
              proof (cases "A' = ?C")
                case True
                thus ?thesis 
                  by (metis "1" Lem_Gradn_id_n cong_diff_3 
                      Bet A0 (Gradn A0 D1 (Suc n)) A Bet A0 A A' 
                      Cong A A' A0 A1 D = Gradn A0 D1 (Suc (Suc n)) 
                      between_equality_2 between_identity)
              next
                case False
                hence "A'  ?C" 
                  by blast
                have "?C Out D A'" 
                proof -
                  have "D  ?C"
                    using D  Gradn A0 D1 (Suc n) by auto
                  moreover
                  have "A'  ?C" 
                    using False by blast
                  moreover
                  have "A0  ?C" 
                    using A0  Gradn A0 D1 (Suc n) by auto
                  moreover
                  have "Bet D ?C A0" 
                    using Bet_Gradn_Gradn_Suc 
                      D = Gradn A0 D1 (Suc (Suc n)) between_symmetry by blast
                  ultimately
                  show ?thesis 
                    by (meson Bet A0 (Gradn A0 D1 (Suc n)) A Bet A0 A A' 
                        between_exchange4 between_symmetry l6_3_2)
                qed
                moreover
                have "?C D Le ?C A'" 
                proof -
                  have "?C D Le A A'" 
                    by (metis "1" Cong_Gradn_Gradn_Suc cong_symmetry 
                        Cong A A' A0 A1 D = Gradn A0 D1 (Suc (Suc n)) 
                        bet__le1213 l5_6)
                  moreover
                  have "A A' Le ?C A'" 
                    by (metis between_exchange3 
                        Bet A0 (Gradn A0 D1 (Suc n)) A 
                        Bet A0 A A' bet__le2313)
                  ultimately
                  show ?thesis 
                    using le_transitivity by blast
                qed
                ultimately
                show ?thesis
                  using between_symmetry l6_13_1 by blast
              qed
            qed
          qed
        qed
        moreover
        have "Cong ?C A' A' ?C" 
          by (simp add: cong_pseudo_reflexivity)
        moreover
        have "Cong D' A' D ?C" 
        proof -
          have "Cong A0 D1 (Gradn A0 D1 (Suc n)) (Gradn A0 D1 (Suc (Suc n)))" 
            using Cong_Gradn_Gradn_Suc by blast
          hence "Cong A0 D1 ?C D" 
            using D = Gradn A0 D1 (Suc (Suc n)) by blast
          hence "Cong A0 D1 D ?C" 
            using not_cong_1243 by blast
          have "Bet A1 D1 A0" 
            by (simp add: "1" between_symmetry)
          have "Cong A1 D1 A D'" 
            using "2" Cong E0 E1 A D' cong_inner_transitivity 
            by blast
          have "Cong A0 A1 A A'" 
            using Cong A A' A0 A1 cong_symmetry by presburger
          hence "Cong A0 D1 D' A'" 
            using  Bet A D' A' Bet A1 D1 A0 Cong A1 D1 A D' 
              l4_3_1 not_cong_2134 by blast
          thus ?thesis 
            using Cong A0 D1 D (Gradn A0 D1 (Suc n)) 
              cong_inner_transitivity by blast
        qed
        ultimately
        show ?thesis 
          using l4_3 by blast
      qed
      ultimately
      show ?thesis 
        using Cong_cases cong_inner_transitivity by blast
    qed
    moreover
    have "Bet A0 D A'" 
    proof -
      have "Bet A0 ?C A'" 
        using Bet A0 (Gradn A0 D1 (Suc n)) A Bet A0 A A' 
          between_exchange4 by blast
      have "Bet ?C D A'" 
      proof -
        have "Grad A0 A1 A" 
          using Grad2 A0 A1 A E0 E1 E grad2__grad123 by blast
        hence "Bet A0 A1 A" 
          by (simp add: grad__bet)
        have "Bet A0 D1 ?C" 
          using Diff__Bet_Gradn_Suc not_bet_distincts by blast
        show ?thesis 
        proof (cases "A0 = ?C")
          case True
          thus ?thesis 
            by (metis Gradn_aux_1_0 Lem_Gradn_id_n 
                Bet A0 (Gradn A0 D1 (Suc n)) A' 
                D = Gradn A0 D1 (Suc (Suc n)))
        next
          case False
          hence "A0  ?C" 
            by blast
          show ?thesis 
          proof (cases "D = ?C")
            case True
            thus ?thesis 
              using between_trivial2 by auto
          next
            case False
            hence "D  ?C"
              by blast
            show ?thesis 
            proof (cases "A' = ?C")
              case True
              thus ?thesis 
                by (metis "1" Lem_Gradn_id_n cong_diff_3
                    Bet A0 (Gradn A0 D1 (Suc n)) A Bet A0 A A' 
                    Cong A A' A0 A1 D = Gradn A0 D1 (Suc (Suc n)) 
                    between_equality_2 between_identity)
            next
              case False
              hence "A'  ?C" 
                by blast
              have "?C Out D A'" 
              proof -
                have "D  ?C"
                  using D  Gradn A0 D1 (Suc n) by auto
                moreover
                have "A'  ?C" 
                  using False by blast
                moreover
                have "A0  ?C" 
                  using A0  Gradn A0 D1 (Suc n) by auto
                moreover
                have "Bet D ?C A0" 
                  using Bet_Gradn_Gradn_Suc 
                    D = Gradn A0 D1 (Suc (Suc n)) 
                    between_symmetry by blast
                ultimately
                show ?thesis 
                  by (meson Bet A0 (Gradn A0 D1 (Suc n)) A'
                      between_symmetry l6_3_2)
              qed
              moreover
              have "?C D Le ?C A'" 
              proof -
                have "?C D Le A A'" 
                  by (metis "1" Cong_Gradn_Gradn_Suc cong_symmetry
                      Cong A A' A0 A1 D = Gradn A0 D1 (Suc (Suc n)) 
                      bet__le1213 l5_6)
                moreover
                have "A A' Le ?C A'" 
                  by (metis between_exchange3 Bet A0 (Gradn A0 D1 (Suc n)) A 
                      Bet A0 A A' bet__le2313)
                ultimately
                show ?thesis 
                  using le_transitivity by blast
              qed
              ultimately
              show ?thesis 
                using l6_13_1 by blast
            qed
          qed
        qed
      qed
      thus ?thesis
        using Bet A0 (Gradn A0 D1 (Suc n)) A' between_exchange2 by blast
    qed
    ultimately
    have " A E. (Grad2 A0 A1 A E0 E1 E  Cong E0 E A D  Bet A0 D A)" 
      by blast
  }
  thus ?thesis by blast
qed

lemma t22_18_aux0:
  shows " A0 D1 A1 E0 E1 D. 
  (Bet A0 D1 A1  Cong E0 E1 A1 D1  D = Gradn A0 D1 (Suc n)) 
  ( A E. (Grad2 A0 A1 A E0 E1 E  Cong E0 E A D  Bet A0 D A))" 
proof (induction n)
  case 0
  show ?case 
    using t22_18_aux_0 by auto
next
  case (Suc n)
  show ?case
    using t22_18_aux_n Suc.IH by blast
qed

(** For every m, there exists n such that A0Dm = A0An - E0En = n(A0A1 - E0E1) (m=n) *)
lemma t22_18_aux1:
  assumes "Bet A0 D1 A1" and
    "Cong E0 E1 A1 D1" and
    "Grad A0 D1 D"
  shows " A E. (Grad2 A0 A1 A E0 E1 E  Cong E0 E A D  Bet A0 D A)"
proof -
  obtain n  where "n  0" and "D = Gradn A0 D1 n" 
    using Grad_def assms(3) by blast
  hence "n  0" 
    by blast
  then obtain m where "n = Suc m" 
    using not0_implies_Suc by presburger
  hence "D = Gradn A0 D1 (Suc m)" 
    by (simp add: D = Gradn A0 D1 n)
  thus ?thesis 
    using assms(1) assms(2) t22_18_aux0 by blast  
qed

lemma t22_18_aux2_0:
  assumes "Saccheri A0 B0 B1 A1" and
    "A = Gradn A0 A1 (Suc 0)" and
    "E = Gradn B0 B1 (Suc 0)" and
    "Saccheri A0 B0 B A"
  shows "B0 B Le B0 E" 
proof -
  have "Per B0 A0 A1" 
    using Saccheri_def assms(1) by blast
  have "Per A0 A1 B1" 
    using Saccheri_def assms(1) by blast
  have "Cong A0 B0 B1 A1" 
    using Saccheri_def assms(1) by blast
  have "A0 A1 OS B0 B1"   
    using Saccheri_def assms(1) by blast
  have "Per B0 A0 A"
    using Saccheri_def assms(4) by blast
  have "Per A0 A B"
    using Saccheri_def assms(4) by blast
  have "A0 A OS B0 B"
    using Saccheri_def assms(4) by blast
  have "Cong A0 B0 B A"
    using Saccheri_def assms(4) by blast
  have "A = A1"
    by (simp add: assms(2))
  have "E = B1"
    by (simp add: assms(3))
  have "Saccheri A0 B0 B A1" 
    using A = A1 assms(4) by auto
  hence "Cong A0 B0 B A1"
    using Saccheri_def by blast
  hence "B = B1" 
  proof-
    have "A0  A1" 
      using A0 A1 OS B0 B1 os_distincts by blast
    have "A1 Out B B1" 
    proof -
      have "Col A1 B B1" 
      proof -
        have "Coplanar A0 B B1 A1" 
        proof -
          have "¬ Col B0 A0 A1" 
            by (meson A0 A1 OS B0 B1 col123__nos not_col_permutation_2)
          moreover
          have "Coplanar B0 A0 A1 B" 
            using A = A1 assms(4) coplanar_perm_7 sac__coplanar by blast
          moreover
          have "Coplanar B0 A0 A1 B1" 
            by (meson A0 A1 OS B0 B1 ncoplanar_perm_8 os__coplanar)
          ultimately
          show ?thesis 
            by (meson coplanar_pseudo_trans ncop_distincts)
        qed
        moreover
        have "A0  A1" 
          using A0  A1 by auto
        moreover
        have "Per B A1 A0" 
          using A = A1 Per A0 A B l8_2 by blast
        moreover
        have "Per B1 A1 A0" 
          using Per_cases Per A0 A1 B1 by blast
        ultimately
        show ?thesis 
          using col_permutation_2 cop_per2__col by blast
      qed
      moreover
      have "A1 A0 OS B B1" 
      proof -
        have "A1 A0 OS B B0" 
          using Saccheri_def A = A1 assms(4) invert_one_side 
            one_side_symmetry by blast
        moreover
        have "A1 A0 OS B0 B1" 
          by (simp add: A0 A1 OS B0 B1 invert_one_side)
        ultimately
        show ?thesis 
          using one_side_transitivity by blast
      qed
      ultimately
      show ?thesis 
        using col_one_side_out by auto
    qed
    moreover
    have "Cong A1 B A0 B0" 
      using Cong_cases Cong A0 B0 B A1 by blast
    moreover
    have "A1 Out B1 B1" 
      using calculation(1) out_diff2 out_trivial by auto
    moreover
    have "Cong A1 B1 A0 B0" 
      using Cong A0 B0 B1 A1 not_cong_3421 by blast
    ultimately
    show ?thesis 
      using l6_11_uniqueness by blast
  qed
  thus ?thesis 
    using E = B1 local.le_cases by blast
qed

lemma t22_18_aux2_Sucn:
  assumes " A A0 A1 B B0 B1 E.
  Saccheri A0 B0 B1 A1  A = Gradn A0 A1 (Suc n) 
  E = Gradn B0 B1 (Suc n)  Saccheri A0 B0 B A  B0 B Le B0 E" 
  shows " A A0 A1 B B0 B1 E.
  Saccheri A0 B0 B1 A1  A = Gradn A0 A1 (Suc(Suc n)) 
  E = Gradn B0 B1 (Suc(Suc n))  Saccheri A0 B0 B A  B0 B Le B0 E" 
proof -
  {
    fix A A0 A1 B B0 B1 E
    assume "Saccheri A0 B0 B1 A1" and
      "A = Gradn A0 A1 (Suc(Suc n))" and
      "E = Gradn B0 B1 (Suc(Suc n))" and
      "Saccheri A0 B0 B A"
    have "A0  B0" 
      using Saccheri A0 B0 B1 A1 sac_distincts by blast
    let ?A = "Gradn A0 A1 (Suc n)"
    let ?E = "Gradn B0 B1 (Suc n)"
    have "Saccheri A0 B0 B ?A  B0 B Le B0 ?E"
      using Saccheri A0 B0 B1 A1 assms by blast
    {
      assume "A0 = ?A"
      have "Bet A0 A1 ?A" 
        using Diff__Bet_Gradn_Suc not_bet_distincts by blast
      hence "A1 = A0" 
        by (metis A0 = Gradn A0 A1 (Suc n) bet_neq32__neq)
      hence False 
        using sac_distincts Saccheri A0 B0 B1 A1 by blast
    }
    hence "A0  ?A" 
      by blast
    have " B. Saccheri A0 B0 B ?A" 
    proof -
      have "Grad A0 A1 ?A" 
        using Grad_def by blast
      hence "Bet A0 A1 ?A" 
        by (simp add: grad__bet)
      have " P. A0 A1 Perp P ?A  A0 A1 OS B0 P"
      proof -
        have "Grad A0 A1 ?A" 
          using Grad_def by blast
        hence "Bet A0 A1 ?A" 
          by (simp add: grad__bet)
        hence "Col A0 A1 ?A" 
          using Col_def by blast
        moreover
        have "¬ Col A0 A1 B0" 
          by (meson Saccheri A0 B0 B1 A1 Col_cases col_trivial_3 
              l8_16_1 sac__perp1214)
        ultimately
        show ?thesis 
          using l10_15 by blast
      qed
      then obtain P where "A0 A1 Perp P ?A" and "A0 A1 OS B0 P" 
        by blast
      hence "P  ?A" 
        using perp_not_eq_2 by blast
      hence " B. ?A Out B P  Cong ?A B A0 B0" 
        using l6_11_existence A0  B0 by simp
      then obtain B' where "?A Out B' P" and "Cong ?A B' A0 B0"
        by blast
      have "Saccheri A0 B0 B' ?A" 
      proof -
        have "Per B0 A0 ?A" 
        proof -
          have "Per B0 A0 A1" 
            by (meson perp_per_2 Saccheri A0 B0 B1 A1 sac__perp1214)
          moreover
          have "Col A0 A1 ?A" 
            using Col_def Bet A0 A1 (Gradn A0 A1 (Suc n)) by blast
          ultimately
          show ?thesis 
            using Saccheri A0 B0 B1 A1 per_col sac_distincts by blast
        qed
        moreover
        have "Per A0 ?A B'"
        proof -
          have "?A A0 Perp ?A P" 
            using Perp_cases A0 A1 Perp P (Gradn A0 A1 (Suc n))
              A0  Gradn A0 A1 (Suc n) Bet A0 A1 (Gradn A0 A1 (Suc n))
              bet_col perp_col1 by blast
          moreover
          have "Col ?A P B'" 
            using Gradn A0 A1 (Suc n) Out B' P col_permutation_5 out_col by blast
          ultimately
          show ?thesis 
            by (meson Perp_cases col_trivial_3 l8_16_1)
        qed
        moreover
        have "Col A0 A1 ?A" 
          using Col_def Bet A0 A1 (Gradn A0 A1 (Suc n)) by blast
        hence "A0 ?A OS B0 P" 
          using A0 A1 OS B0 P A0  Gradn A0 A1 (Suc n) 
            col_one_side by blast
        hence "?A A0 OS B0 P" 
          using invert_one_side by blast
        hence "?A A0 OS B0 B'" 
          using Gradn A0 A1 (Suc n) Out B' P l6_6
            out_out_one_side by blast
        hence "A0 ?A OS B0 B'" 
          using invert_one_side by blast
        ultimately
        show ?thesis 
          using Saccheri_def Cong (Gradn A0 A1 (Suc n)) B' A0 B0
            not_cong_4312 by blast
      qed
      thus ?thesis 
        by blast
    qed
    then obtain B' where "Saccheri A0 B0 B' ?A" 
      by blast
    hence "B0 B' Le B0 ?E" 
      using Saccheri A0 B0 B1 A1 assms by blast
    have "Per B0 A0 ?A"
      using Saccheri A0 B0 B' ?A Saccheri_def by blast
    have "Per A0 ?A B'"
      using Saccheri A0 B0 B' ?A Saccheri_def by blast
    have "Cong A0 B0 B' ?A"
      using Saccheri A0 B0 B' ?A Saccheri_def by blast
    have "A0 ?A OS B0 B'"
      using Saccheri A0 B0 B' ?A Saccheri_def by blast
    obtain C where "Bet B0 B' C" and "Cong B' C ?E E" 
      using segment_construction by blast
    have "Cong B0 B1 B' B" 
    proof -
      have "Saccheri A0 B0 B1 A1" 
        by (simp add: Saccheri A0 B0 B1 A1)
      moreover
      have "Saccheri ?A B' B A" 
      proof -
        have "Saccheri A0 B0 B' ?A" 
          using Saccheri A0 B0 B' (Gradn A0 A1 (Suc n)) by auto
        moreover
        have "?A  A" 
          by (metis Gradn_uniq_aux_1 A = Gradn A0 A1 (Suc (Suc n)) 
              A0  Gradn A0 A1 (Suc n) grad_rec_a_a)
        moreover
        have "Coplanar A0 B0 ?A A" 
          by (metis Bet_Gradn_Gradn_Suc ncop__ncols 
              A = Gradn A0 A1 (Suc (Suc n)) bet_col)
        ultimately
        show ?thesis 
          by (meson cop_sac2__sac Saccheri A0 B0 B A)
      qed
      moreover
      have "Cong A0 B0 ?A B'" 
        using Cong A0 B0 B' (Gradn A0 A1 (Suc n)) 
          cong_right_commutativity by blast
      moreover
      have "Cong A0 A1 ?A A" 
        using Cong_Gradn_Gradn_Suc A = Gradn A0 A1 (Suc (Suc n)) by blast
      ultimately
      show ?thesis 
        using cong2_sac2__cong by blast
    qed
    have "Cong B B' B' C" 
    proof -
      have "Cong B B' ?E E" 
        using Cong_Gradn_Gradn_Suc 
          Cong B0 B1 B' B E = Gradn B0 B1 (Suc (Suc n)) 
          cong_transitivity not_cong_3421 by blast
      moreover
      have "Cong ?E E B' C" 
        using Cong B' C (Gradn B0 B1 (Suc n)) E 
          cong_inner_transitivity cong_reflexivity by blast
      ultimately
      show ?thesis 
        by (meson cong_transitivity)
    qed
    hence "B0 B Le B0 C"
      by (meson  Bet B0 B' C Cong_cases cong_reflexivity 
          triangle_inequality_2)
    moreover
    have "B0 C Le B0 E" 
    proof -
      have "Bet B0 B' C" 
        by (simp add: Bet B0 B' C)
      moreover
      have "Bet B0 ?E E" 
        using Bet_Gradn_Gradn_Suc E = Gradn B0 B1 (Suc (Suc n)) by blast
      moreover
      have "B0 B' Le B0 ?E" 
        using B0 B' Le B0 (Gradn B0 B1 (Suc n)) by auto
      moreover
      have "B' C Le ?E E" 
        using Cong B' C (Gradn B0 B1 (Suc n)) E cong__le by blast
      ultimately
      show ?thesis 
        using Cong B' C (Gradn B0 B1 (Suc n)) E 
          bet2_le2__le1346 cong__le by blast
    qed
    ultimately
    have "B0 B Le B0 E"
      using le_transitivity by blast
  }
  thus ?thesis 
    by blast
qed

lemma t22_18_aux2_n:
  shows " A A0 A1 B B0 B1 E.
  Saccheri A0 B0 B1 A1  A = Gradn A0 A1 (Suc n) 
  E = Gradn B0 B1 (Suc n)  Saccheri A0 B0 B A 
   B0 B Le B0 E" 
proof (induction n)
  case 0
  show ?case
    using t22_18_aux2_0 by blast
next
  case (Suc n)
  show ?case
    using t22_18_aux2_Sucn Suc.IH by blast
qed

(** For every n, B0Bn is lower than or equal to n times B0B1 *)
lemma t22_18_aux2:
  assumes "Saccheri A0 B0 B1 A1" and
    "Grad2 A0 A1 A B0 B1 E" and
    "Saccheri A0 B0 B A"
  shows "B0 B Le B0 E" 
proof -
  obtain n where "n  0" and "A = Gradn A0 A1 n" and 
    "E = Gradn B0 B1 n" 
    using Grad2_def assms(2) by blast
  obtain m where "n = Suc m" 
    using n  0 not0_implies_Suc by blast
  hence "A = Gradn A0 A1 (Suc m)  E = Gradn B0 B1 (Suc m)" 
    using A = Gradn A0 A1 n E = Gradn B0 B1 n by blast
  show ?thesis 
    using t22_18_aux2_n 
      A = Gradn A0 A1 (Suc m)  E = Gradn B0 B1 (Suc m) 
      assms(1) assms(3) by blast
qed

lemma t22_18:
  assumes "archimedes_axiom"
  shows " A0 B0 B1 A1. Saccheri A0 B0 B1 A1  ¬ (B0 B1 Lt A1 A0)"
proof -
  {
    fix A0 B0 B1 A1
    assume "Saccheri A0 B0 B1 A1"
    assume "B0 B1 Lt A1 A0"
    hence "B0 B1 Le A1 A0"
      using cong__nlt lt__le by blast
    then obtain D1 where "Bet A1 D1 A0" and "Cong B0 B1 A1 D1" 
      using Le_def by blast
    obtain C0 where "Bet A0 D1 C0" and "Cong D1 C0 A0 B0" 
      using segment_construction by fastforce
    obtain C where "Bet A0 C0 C" and "Cong C0 C A0 B0" 
      using segment_construction by fastforce
    {
      fix D
      assume "Grad A0 D1 D"
      have "A0 D Lt A0 C" 
      proof (cases "A0 = D1")
        case True
        thus ?thesis 
          using B0 B1 Lt A1 A0 Cong B0 B1 A1 D1 cong__nlt by blast
      next 
        case False
        obtain A E where "Grad2 A0 A1 A B0 B1 E  Cong B0 E A D  Bet A0 D A" 
          using t22_18_aux1 Bet_cases Bet A1 D1 A0 
            Cong B0 B1 A1 D1 Grad A0 D1 D by blast
        have "Grad2 A0 A1 A B0 B1 E" 
          using Grad2 A0 A1 A B0 B1 E  Cong B0 E A D  Bet A0 D A by blast
        hence "Grad A0 A1 A" 
          using grad2__grad123 by blast
        hence "Bet A0 A1 A" 
          by (simp add: grad__bet)
        have "Cong B0 E A D" 
          using Grad2 A0 A1 A B0 B1 E  Cong B0 E A D  Bet A0 D A by blast
        have "Bet A0 D A" 
          using Grad2 A0 A1 A B0 B1 E  Cong B0 E A D  Bet A0 D A by blast
        have " P. A0 A1 Perp A P  A0 A1 OS B0 P"
        proof -
          have "Col A0 A1 A" 
            by (simp add: Col_def Bet A0 A1 A)
          moreover
          have "A0 A1 ParStrict B0 B1" 
            by (simp add: Saccheri A0 B0 B1 A1 sac__pars1423)
          hence "¬ Col A0 A1 B0"
            using par_strict_not_col_1 by blast
          ultimately
          show ?thesis 
            using Perp_cases l10_15 by blast
        qed
        then obtain P where "A0 A1 Perp A P" and "A0 A1 OS B0 P"
          by blast
        have "P  A" 
          using A0 A1 Perp A P perp_not_eq_2 by blast
        have "A0  B0" 
          using A0 A1 OS B0 P os_distincts by blast
        then obtain B where "A Out B P" and "Cong A B A0 B0"
          using l6_11_existence P  A by blast
        have "Saccheri A0 B0 B A" 
        proof -
          have "Per B0 A0 A" 
          proof -
            have "A0  A1" 
              using A0 A1 OS B0 P os_distincts by blast
            moreover
            have "Per B0 A0 A1" 
              by (meson perp_per_2 Saccheri A0 B0 B1 A1 sac__perp1214)
            moreover
            have "Col A0 A1 A" 
              by (simp add: Col_def Bet A0 A1 A)
            ultimately
            show ?thesis 
              using per_col by blast
          qed
          moreover
          have "Per A0 A B"
          proof -
            have "A  B" 
              using A Out B P out_distinct by blast
            moreover
            have "A A0 Perp A P" 
            proof -
              have "A0  A" 
                using False Bet A0 D A Grad A0 D1 D 
                  bet_neq32__neq grad_neq__neq13 by blast
              moreover
              have "A0 A1 Perp P A" 
                using Perp_cases A0 A1 Perp A P by blast
              moreover
              have "Col A0 A1 A" 
                by (simp add: Col_def Bet A0 A1 A)
              ultimately
              show ?thesis 
                using A0 A1 Perp A P col_trivial_3 perp_col2 
                by presburger
            qed
            moreover
            have "Col A P B" 
              by (simp add: A Out B P l6_6 out_col)
            ultimately
            show ?thesis 
              by (meson perp_per_2 perp_col1)
          qed
          moreover
          have "Cong A0 B0 B A" 
            using Cong A B A0 B0 not_cong_4312 by blast
          moreover
          have "A A0 OS B0 B"
          proof -
            have "A A0 OS B0 P" 
              using A0 A1 OS B0 P Bet A0 A1 A bet_col 
                bet_neq32__neq col_one_side invert_one_side by blast
            moreover
            have "A Out P B" 
              using Out_cases A Out B P by blast
            ultimately
            show ?thesis 
              using out_out_one_side by blast
          qed
          hence "A0 A OS B0 B" 
            using invert_one_side by blast
          ultimately
          show ?thesis 
            using Saccheri_def by blast
        qed
        have "B0 B Le B0 E"
          using Saccheri A0 B0 B1 A1 Grad2 A0 A1 A B0 B1 E 
            Saccheri A0 B0 B A t22_18_aux2 by blast
        have "B0 E Le A A0" 
          using Le_def Bet A0 D A Cong B0 E A D 
            between_symmetry by blast
        hence "B0 B Le A A0" 
          using le_transitivity B0 B Le B0 E by blast
        then obtain Q where "Bet A Q A0" and "Cong B0 B A Q" 
          using Le_def by blast
        have "A0 D Le A0 Q" 
        proof -
          have "Bet A0 Q A" 
            using Bet_cases Bet A Q A0 by blast
          moreover
          have "Q A Le D A" 
          proof -
            have "Cong B0 B Q A" 
              using Cong_cases Cong B0 B A Q by blast
            moreover
            have "Cong B0 E D A" 
              using Cong_cases Cong B0 E A D by blast
            ultimately
            show ?thesis 
              using B0 B Le B0 E l5_6 by blast
          qed
          moreover
          have "A0 A Le A0 A" 
            using between_trivial2 l5_12_a by auto
          ultimately
          show ?thesis 
            using Bet A0 D A bet2_le2__le1245 by blast
        qed
        moreover
        have "A  A0" 
          using False Bet A0 D A Grad A0 D1 D between_identity 
            grad_neq__neq13 by blast
        then obtain B0' where "A0 Out B0' A  Cong A0 B0' A0 B0"
          using A0  B0 l6_11_existence by presburger
        have "A0 Out B0' A"
          using A0 Out B0' A  Cong A0 B0' A0 B0 by blast
        have "Cong A0 B0' A0 B0"
          using A0 Out B0' A  Cong A0 B0' A0 B0 by blast
        obtain B' where "Bet A0 B0' B'" and "Cong B0' B' B0 B" 
          using segment_construction by blast
        obtain A' where "Bet B0' B' A'" and "Cong B' A' B A" 
          using segment_construction by blast
        have "A0 A Le A0 A'" 
        proof -
          obtain B'' where "Bet A0 B0' B''" and "Cong B0' B'' B0 A"
            using segment_construction by blast
          have "A0 A Le A0 B''" 
            by (meson Bet A0 B0' B'' Cong A0 B0' A0 B0 
                Cong B0' B'' B0 A cong__le3412 l2_11_b lt__le not_cong_3412 
                triangle_strict_inequality_2)
          moreover
          have "A0 B'' Le A0 A'" 
          proof -
            have "B0'  B'" 
              using sac_distincts 
              using Cong B0' B' B0 B Saccheri A0 B0 B A 
                cong_reverse_identity by blast
            hence "Bet A0 B0' A'" 
              using  Bet B0' B' A' Bet A0 B0' B' outer_transitivity_between 
              by blast 
            moreover
            have "A0 B0' Le A0 B0'" 
              using local.le_cases by auto 
            moreover
            have "B0' B'' Le B0' A'" 
            proof -
              have "B0 A Le B0' A'" 
              proof -
                have "Cong B0 B B0' B'" 
                  using Cong B0' B' B0 B not_cong_3412 by blast
                moreover
                have "Cong B A B' A'" 
                  using Cong B' A' B A cong_symmetry by presburger
                ultimately
                show ?thesis 
                  using Bet B0' B' A' triangle_inequality_2 by blast
              qed
              moreover
              have "Cong B0 A B0' B''" 
                using Cong B0' B'' B0 A cong_symmetry by blast
              moreover
              have "Cong B0' A' B0' A'" 
                by (simp add: cong_reflexivity)
              ultimately
              show ?thesis 
                using l5_6 by blast
            qed
            ultimately
            show ?thesis 
              using Bet A0 B0' B'' bet2_le2__le1346 by blast
          qed
          ultimately
          show ?thesis 
            using le_transitivity by blast
        qed
        have "B0 B Le A' B0'" 
        proof -
          have "B0' B' Le B0' A'" 
            by (simp add: Bet B0' B' A' l5_12_a)
          moreover
          have "Cong B0' A' A' B0'" 
            by (simp add: cong_pseudo_reflexivity)
          ultimately
          show ?thesis 
            using Cong B0' B' B0 B l5_6 by blast
        qed
        obtain Q' where "Bet A' Q' B0'" and "Cong B0 B A' Q'" 
          using Le_def B0 B Le A' B0' by blast
        have "B0'  B'" 
          using sac_distincts Cong B0' B' B0 B Saccheri A0 B0 B A 
            cong_reverse_identity by blast
        hence "Bet A0 B0' A'"
          using  Bet A0 B0' B' Bet B0' B' A' 
            outer_transitivity_between by blast
        have "A0 Q Lt A0 C" 
        proof -
          have "A0 Q Le A0 Q'" 
          proof -
            have "Bet A0 Q' A'" 
              using Bet A' Q' B0' Bet A0 B0' A' bet3__bet 
                between_trivial by blast
            moreover
            have "Bet A0 Q A" 
              using Bet_cases Bet A Q A0 by auto
            moreover
            have "Q' A' Le Q A" 
              by (meson le_reflexivity Cong B0 B A Q 
                  Cong B0 B A' Q' l5_6 le_comm)
            ultimately
            show ?thesis 
              using A0 A Le A0 A' bet2_le2__le1245 by blast
          qed
          moreover
          have "A0 Q' Lt A0 C" 
          proof -
            have "Bet A0 D1 C" 
              using Bet A0 C0 C Bet A0 D1 C0 between_exchange4 by blast
            hence "D1 C Lt A0 C" 
              by (simp add: False bet__lt2313)
            moreover
            have "Cong D1 C A0 Q'" 
            proof -
              have "Bet D1 C0 C"
                using Bet A0 C0 C Bet A0 D1 C0 between_exchange3 by blast
              moreover
              have "Bet A0 B0' Q'" 
                using Bet A' Q' B0' Bet A0 B0' A' between_exchange3 
                  between_symmetry by blast
              moreover
              have "Cong D1 C0 A0 B0'" 
                by (metis Cong_cases Cong A0 B0' A0 B0 
                    Cong D1 C0 A0 B0 cong_inner_transitivity)
              moreover
              have "Cong B0' Q' A B" 
              proof -
                have "Cong B0' B' A' Q'" 
                  using Cong B0 B A' Q' Cong B0' B' B0 B 
                    cong_transitivity by blast
                have "Cong B0' Q' A' B'" 
                proof (cases "Bet B0' Q' B'")
                  case True
                  hence "Bet A' B' Q'" 
                    using Bet_perm Bet B0' B' A' between_exchange3 by blast
                  moreover
                  have "Cong Q' B' B' Q'" 
                    by (simp add: cong_pseudo_reflexivity)
                  ultimately
                  show ?thesis 
                    using Cong B0' B' A' Q' True l4_3 by blast
                next
                  case False
                  hence "Q'  B0'" 
                    using between_trivial2 by blast
                  hence "A'  B0'" 
                    using Bet A' Q' B0' between_identity by blast
                  have "B0' Out B' Q'" 
                  proof -
                    have "B0' Out B' A'" 
                      using B0'  B' Bet B0' B' A' bet_out by force
                    moreover
                    have "B0' Out A' Q'" 
                      by (simp add: Bet A' Q' B0' Q'  B0' bet_out_1 l6_6)
                    ultimately
                    show ?thesis 
                      using l6_7 by blast
                  qed
                  hence "Bet B0' B' Q'" 
                    using False Out_def by blast
                  moreover
                  have "Bet A' Q' B'" 
                    using Bet A' Q' B0' between_exchange3 
                      between_symmetry calculation by blast
                  moreover
                  have "Cong B' Q' Q' B'"
                    by (simp add: cong_pseudo_reflexivity)
                  ultimately
                  show ?thesis 
                    using Cong B0' B' A' Q' l2_11_b by blast
                qed
                moreover
                have "Cong A' B' A B" 
                  using Cong B' A' B A not_cong_2143 by blast
                ultimately
                show ?thesis 
                  using cong_transitivity by blast
              qed
              hence "Cong B0' Q' A0 B0" 
                using  Cong A B A0 B0 cong_transitivity by blast
              hence "Cong C0 C B0' Q'" 
                using cong_inner_transitivity Cong C0 C A0 B0 cong_symmetry by blast
              ultimately
              show ?thesis 
                using l2_11_b by blast
            qed
            moreover
            have "Cong A0 C A0 C" 
              by (simp add: cong_reflexivity)
            ultimately
            show ?thesis 
              using cong2_lt__lt by blast
          qed
          ultimately
          show ?thesis 
            by (meson le1234_lt__lt)
        qed
        ultimately
        show ?thesis 
          using le1234_lt__lt by blast
      qed
    }
    hence " D. Grad A0 D1 D  A0 D Lt A0 C" 
      by blast
    have "¬ Cong B0 B1 A1 A0"
      using B0 B1 Lt A1 A0 cong__nlt lt__le by blast
    have"A0  D1  Reach A0 D1 A0 C" 
      using archimedes_axiom_def assms by blast
    then obtain D where "Grad A0 D1 D  A0 C Le A0 D" 
      using Reach_def Cong B0 B1 A1 D1 ¬ Cong B0 B1 A1 A0 by blast
    hence "A0 D Lt A0 C" 
      by (simp add: D. Grad A0 D1 D  A0 D Lt A0 C)
    moreover
    have "A0 C Lt A0 C" 
      using Grad A0 D1 D  A0 C Le A0 D calculation lt__nle by blast  
    ultimately
    have False 
      using nlt by auto
  }
  thus ?thesis 
    by blast
qed

lemma t22_19:
  assumes "archimedes_axiom" 
  shows " A B C D. Saccheri A B C D  ¬ Obtuse A B C" 
proof -
  {
    fix A B C D
    assume "Saccheri A B C D"
    hence "¬ C B Lt A D" 
      using assms lt_comm t22_18 by blast
    moreover
    assume "Obtuse A B C"
    hence "C B Lt A D" 
      using Saccheri A B C D lt_left_comm lt_sac__obtuse_aux2 by blast
    ultimately
    have False 
      by blast
  }
  thus ?thesis 
    by blast
qed

lemma archi__obtuse_case_elimination:
  assumes "archimedes_axiom"
  shows "¬ hypothesis_of_obtuse_saccheri_quadrilaterals" 
proof -
  have "¬ ( A B C D. Saccheri A B C D  Obtuse A B C)"
    using assms ex_saccheri t22_19 by blast
  thus ?thesis 
    using hypothesis_of_obtuse_saccheri_quadrilaterals_def by blast
qed

lemma t22_23_aux:
  assumes "¬ Col A M N" and
    "Per B C A" and
    "A  C" and
    "M Midpoint A B" and
    "Per M N A" and
    "Col A C N" and
    "M Midpoint N L"
  shows "Bet A N C  Lambert N L B C  Cong B L A N"
proof -
  have "A  M" 
    using assms(1) col_trivial_1 by blast
  have "N  M" 
    using assms(1) col_trivial_2 by blast
  have "A  N" 
    using assms(1) col_trivial_3 by blast
  have "L  N" 
    using N  M assms(7) l7_3 by blast
  have "A  B" 
    using A  M assms(4) l7_3 by blast
  hence "B  M" 
    using assms(4) is_midpoint_id_2 by blast
  have "Bet A N C" 
  proof -
    have "Bet A M B" 
      by (simp add: assms(4) midpoint_bet)
    moreover
    have "Col A N C" 
      using assms(6) not_col_permutation_5 by blast
    moreover
    have "Per A N M" 
      by (simp add: assms(5) l8_2)
    moreover
    have "Per A C B" 
      by (simp add: assms(2) l8_2)
    ultimately
    show ?thesis 
      by (metis per23_preserves_bet A  N assms(3))
  qed
  moreover
  have "A M N CongA B M L" 
    using A  M N  M assms(4) assms(7) l7_3_2 
      symmetry_preserves_conga by blast
  hence "Cong A N B L  M A N CongA M B L  M N A CongA M L B" 
    using per23_preserves_bet Cong_cases l11_49 A  N assms(4) 
      assms(7) midpoint_cong by meson
  have "Lambert N L B C" 
  proof -
    {
      assume "C = N"
      have "M = B"
      proof -
        have "¬ Col A M N"
          by (simp add: assms(1))
        moreover
        have "N  M" 
          by (simp add: N  M)
        moreover
        have "Col A M M" 
          using not_col_distincts by blast
        moreover
        have "Bet A M B" 
          by (simp add: assms(4) midpoint_bet)
        hence "Col A M B" 
          by (simp add: Col_def)
        moreover
        have "Col N M M" 
          by (simp add: col_trivial_2)
        moreover
        have "Coplanar A B M N" 
          by (meson A  M calculation(3) calculation(4) 
              col__coplanar col_transitivity_1)
        hence "Col N M B" 
          using C = N A  N assms(2) assms(5) col_permutation_3 
            cop_per2__col by blast
        moreover
        show ?thesis 
          using calculation(1) calculation(4) calculation(6) l6_16_1 by blast
      qed
      hence False 
        using B  M by auto
    }
    hence "C  N" 
      by blast
    moreover  
    have "L  B"  
      using A  N Cong A N B L  M A N CongA M B L  M N A CongA M L B 
        cong_diff by blast
    moreover
    have "B  C"  
      using A  B assms(1) assms(4) assms(6) col_permutation_1 
        l6_16_1 midpoint_col by blast
    moreover
    have "N  L" 
      using L  N by auto
    moreover
    have "Per L N C"  
      by (metis per_col A  N N  M assms(5) assms(6) assms(7) 
          col_per2__per l8_20_1_R1 midpoint_col not_col_permutation_1)
    moreover
    have "Per N C B"  
      by (metis per_col assms(2) assms(3) assms(6) col_permutation_4 l8_2)
    moreover
    have "Per N L B"  
    proof -
      have "Per M N A" 
        by (simp add: assms(5))
      moreover
      have "M N A CongA N L B" 
      proof -
        have "M N A CongA M L B" 
          by (simp add: Cong A N B L  M A N CongA M B L  M N A CongA M L B)
        moreover
        have "N Out M M" 
          using N  M out_trivial by auto
        moreover
        have "N Out A A" 
          using A  N Bet A N C C  N l6_3_2 by blast
        moreover
        have "M  L" 
          using N  L assms(7) is_midpoint_id_2 by blast
        have "Bet L M N" 
          using assms(7) Bet_perm Midpoint_def by blast
        hence "L Out N M" 
          by (simp add: Out_def M  L N  L)
        moreover
        have "L Out B B" 
          using L  B out_trivial by auto
        ultimately
        show ?thesis 
          using l11_10 by blast
      qed
      ultimately
      show ?thesis 
        using l11_17 by blast
    qed
    moreover
    have "Coplanar N L B C" 
    proof -
      have "Coplanar B C N M" 
      proof -
        have "Coplanar C N A B" 
          by (simp add: Bet A N C bet__coplanar between_symmetry)
        moreover
        have "Bet A M B" 
          by (simp add: assms(4) midpoint_bet)
        hence "Col B A M" 
          using Col_def by auto
        ultimately
        show ?thesis 
          using A  B col_cop__cop coplanar_perm_14 
            coplanar_perm_7 by blast
      qed
      moreover
      have "Bet N M L" 
        using assms(7) midpoint_bet by blast
      hence "Col N M L" 
        by (simp add: Col_def)
      ultimately
      show ?thesis 
        using N  M col_cop__cop ncoplanar_perm_16 by blast
    qed
    ultimately
    show ?thesis 
      using Lambert_def by blast
  qed
  moreover
  have "Cong A N B L"
    using Cong A N B L  M A N CongA M B L  M N A CongA M L B by blast
  hence "Cong B L A N" 
    by (simp add: cong_symmetry)
  ultimately
  show ?thesis 
    by simp
qed

lemma t22_23:
  assumes "¬ hypothesis_of_obtuse_saccheri_quadrilaterals"
  shows " A B C M N L.
  ¬ Col A M N  Per B C A  A  C  M Midpoint A B 
  Per M N A  Col A C N  M Midpoint N L 
  (Bet A N C  N C Le A N  L N Le B C)"
proof -
  {
    fix A B C M N L
    assume "¬ Col A M N" and 
      "Per B C A" and
      "A  C" and
      "M Midpoint A B" and
      "Per M N A" and
      "Col A C N" and
      "M Midpoint N L" 
    hence "Bet A N C  Lambert N L B C  Cong B L A N" 
      using t22_23_aux by blast
    have "Bet A N C" 
      using Bet A N C  Lambert N L B C  Cong B L A N by blast
    moreover
    have "Lambert N L B C" 
      using Bet A N C  Lambert N L B C  Cong B L A N by blast
    have "Cong B L A N" 
      using Bet A N C  Lambert N L B C  Cong B L A N by blast
    have "¬ Obtuse L B C" 
      using Bet A N C  Lambert N L B C  Cong B L A N assms 
        lam_obtuse__oah_1 by blast
    have "N L OS B C"
      using Lambert N L B C by (simp add: lam__os)
    have "Lambert N C B L" 
      by (simp add: lam_perm  Lambert N L B C)
    hence "N C OS B L"
      by (simp add: lam__os)
    have "N C Le A N  L N Le B C" 
    proof (cases "Acute L B C")
      case True
      have "N C Lt A N" 
      proof -
        have "N C Lt B L" 
          by (simp add: acute_lam__lt True 
              Bet A N C  Lambert N L B C  Cong B L A N lt_right_comm)
        moreover
        have "Cong N C N C" 
          by (simp add: cong_reflexivity)
        moreover
        have "Cong B L A N" 
          by (simp add: Cong B L A N)
        ultimately
        show ?thesis 
          using cong2_lt__lt by blast
      qed
      moreover
      have "N L Lt B C" 
      proof -
        have "Per L N C" 
          by (metis Col_def Bet A N C M Midpoint N L 
              Per M N A ¬ Col A M N between_trivial col_per2__per 
              l8_2 l8_5 midpoint_col)
        moreover
        have "Per N C B" 
          by (metis A  C Col A C N Per B C A
              col_permutation_4 l8_2 per_col)
        moreover
        have "N C OS L B" 
          by (simp add: N C OS B L one_side_symmetry)
        moreover
        have "L B C LtA N L B"
        proof -
          have "Acute L B C" 
            by (simp add: True)
          moreover
          have "Per N L B"
            using Lambert N L B C Lambert_def by blast 
          ultimately
          show ?thesis 
            by (metis Bet A N C  Lambert N L B C  Cong B L A N 
                acute_per__lta lam__os os_distincts)
        qed
        ultimately
        show ?thesis 
          by (simp add: lta_os_per2__lt)
      qed
      hence "L N Lt B C" 
        by (meson lt_left_comm)
      ultimately
      show ?thesis 
        by (simp add: lt__le)
    next
      case False
      hence "Per L B C" 
        by (metis N C OS B L N L OS B C ¬ Obtuse L B C 
            angle_partition os_distincts)
      have "Cong N C B L" 
        using Lambert N L B C Per L B C cong_right_commutativity 
          lam_per__cong by blast
      hence "Cong N C A N"
        using Cong B L A N cong_transitivity by blast
      moreover
      have "Cong L N B C" 
        using Lambert N C B L Per L B C cong_commutativity 
          l8_2 lam_per__cong by blast
      ultimately
      show ?thesis 
        by (simp add: cong__le)
    qed
    ultimately
    have "Bet A N C  N C Le A N  L N Le B C" 
      by blast
  }
  thus ?thesis 
    by blast
qed

lemma t22_24_aux_0_a:
  assumes (*"¬ hypothesis_of_obtuse_saccheri_quadrilaterals" and*)
    "¬ Col A B0 C0" and 
    "A C0 Perp B0 C0" and
    "B = GradExpn A B0 (Suc 0)" and
    "E = GradExpn B0 C0 (Suc 0)" and
    "A C0 Perp B C" and
    "Col A C0 C"
  shows "B0 E Le B C" 
proof -
  have "A  B0" 
    using assms(1) col_trivial_1 by auto
  hence "B = B0"
    using assms(3) by simp
  have "B0  C0" 
    using assms(1) col_trivial_2 by fastforce
  hence "C0 = E"
    using assms(4) by simp
  thus ?thesis 
    using B = B0 assms(2) assms(6) assms(5) col_trivial_2 
      l8_18_uniqueness le_reflexivity by blast
qed

lemma t22_24_aux_0:
  (* assumes "¬ hypothesis_of_obtuse_saccheri_quadrilaterals"*)
  shows " A B0 C0 B C E.
  ¬ Col A B0 C0  A C0 Perp B0 C0  
  (B = GradExpn A B0 (Suc 0))  (E = GradExpn B0 C0 (Suc 0))  
  A C0 Perp B C  Col A C0 C 
  B0 E Le B C"
  using t22_24_aux_0_a by blast

lemma t22_24_aux_suc:
  assumes "¬ hypothesis_of_obtuse_saccheri_quadrilaterals" and
    " A B0 C0 B C E.
  ¬ Col A B0 C0  A C0 Perp B0 C0 
  (B = GradExpn A B0 (Suc n))  (E = GradExpn B0 C0 (Suc n))  
  A C0 Perp B C  Col A C0 C 
  B0 E Le B C" 
  shows
    " A B0 C0 B C E.
  ¬ Col A B0 C0  A C0 Perp B0 C0 
  (B = GradExpn A B0 (Suc(Suc n)))  (E = GradExpn B0 C0 (Suc(Suc n)))  
  A C0 Perp B C  Col A C0 C 
  B0 E Le B C" 
proof -
  {
    fix A B0 C0 B' C' E'
    assume "¬ Col A B0 C0" and 
      "A C0 Perp B0 C0" and
      "B' = GradExpn A B0 (Suc(Suc n))" and
      "E' = GradExpn B0 C0 (Suc(Suc n))" and
      "A C0 Perp B' C'" and
      "Col A C0 C'"
    let ?B = "GradExpn A B0 (Suc n)"
    have "Cong A ?B ?B B'" 
      using GradExpn_4 B' = GradExpn A B0 (Suc (Suc n)) by presburger
    have "Bet A ?B B'" 
      using GradExpn_3 B' = GradExpn A B0 (Suc (Suc n)) by blast
    have "Bet A B0 ?B" 
      using GradExpn_2 by auto
    let ?E = "GradExpn B0 C0 (Suc n)"
    have "Cong B0 ?E ?E E'" 
      using GradExpn_4 E' = GradExpn B0 C0 (Suc (Suc n)) by presburger
    have "Bet B0 ?E E'" 
      using GradExpn_3 E' = GradExpn B0 C0 (Suc (Suc n)) by blast
    have "¬ Col A C0 ?B" 
      by (metis Bet A B0 (GradExpn A B0 (Suc n)) ¬ Col A B0 C0 
          bet_col bet_neq12__neq colx not_col_distincts not_col_permutation_5)
    then obtain C where "Col A C0 C" and "A C0 Perp ?B C" 
      using l8_18_existence by blast
    have "B0 ?E Le ?B C" 
      using A C0 Perp B0 C0 A C0 Perp (GradExpn A B0 (Suc n)) C 
        Col A C0 C ¬ Col A B0 C0 assms(2) by blast
    obtain D where "?B Midpoint C D" 
      using symmetric_point_construction by blast
    have "B0 E' Le D C" 
    proof -
      have "Bet B0 ?E E'" 
        using Bet B0 (GradExpn B0 C0 (Suc n)) E' by blast
      moreover
      have "Bet D ?B C" 
        by (metis midpoint_bet  GradExpn A B0 (Suc n) Midpoint C D 
            between_symmetry)
      moreover
      have "B0 ?E Le D ?B" 
        by (meson B0 (GradExpn B0 C0 (Suc n)) Le (GradExpn A B0 (Suc n)) C 
            GradExpn A B0 (Suc n) Midpoint C D cong_pseudo_reflexivity 
            l5_6 le_right_comm midpoint_cong)
      moreover
      have "?E E' Le ?B C" 
        by (meson Cong B0 (GradExpn B0 C0 (Suc n)) (GradExpn B0 C0 (Suc n)) E' 
            GradExpn A B0 (Suc n) Midpoint C D calculation(3) 
            l5_6 l7_2 midpoint_cong)
      ultimately
      show ?thesis 
        using bet2_le2__le1346 by blast
    qed
    moreover
    {
      assume "A = C"
      have "A = C0" 
      proof -
        have "¬ Col A C0 B0" 
          using ¬ Col A B0 C0 col_permutation_5 by blast
        moreover
        have "Col A C0 A" 
          using col_trivial_3 by blast
        moreover
        have "A C0 Perp B0 A" 
        proof -
          have "A C0 Perp A ?B" 
            using A = C A C0 Perp (GradExpn A B0 (Suc n)) C
              perp_right_comm by blast
          moreover
          have "Col A ?B B0" 
            using Bet A B0 (GradExpn A B0 (Suc n)) bet_col 
              not_col_permutation_5 by blast
          ultimately
          show ?thesis 
            using ¬ Col A C0 B0 not_col_distincts perp_col2_bis by blast
        qed
        moreover
        have "Col A C0 C0" 
          by (simp add: col_trivial_2)
        ultimately
        show ?thesis 
          using A C0 Perp B0 C0 l8_18_uniqueness by blast
      qed
      hence False 
        using ¬ Col A B0 C0 col_trivial_3 by blast
    }
    hence "A  C" 
      by auto
    {
      assume "A = C'"
      have "A = C0" 
      proof -
        have "¬ Col A C0 B0" 
          using ¬ Col A B0 C0 col_permutation_5 by blast
        moreover
        have "Col A C0 A" 
          using col_trivial_3 by blast
        moreover
        have "A C0 Perp B0 A" 
        proof -
          have "A C0 Perp A B'" 
            using A = C' A C0 Perp B' C' perp_right_comm by blast
          moreover
          have "Bet A B0 B'" 
            using B' = GradExpn A B0 (Suc(Suc n)) GradExpn_2 by blast
          hence "Col A B' B0" 
            by (simp add: bet_col col_permutation_5)
          ultimately
          show ?thesis 
            using ¬ Col A C0 B0 not_col_distincts perp_col2_bis by blast
        qed
        moreover
        have "Col A C0 C0" 
          by (simp add: col_trivial_2)
        ultimately
        show ?thesis 
          using A C0 Perp B0 C0 l8_18_uniqueness by blast
      qed
      hence False 
        using ¬ Col A B0 C0 col_trivial_3 by blast
    }
    hence "A  C'"
      by blast
    have "Per A C ?B"  
      using Col A C0 C A C0 Perp (GradExpn A B0 (Suc n)) C 
        col_trivial_3 l8_16_1 l8_2 by blast
    have "D C Le B' C'"
    proof -
      have "¬ Col A ?B C" 
        using NCol_perm A  C Col A C0 C ¬ Col A C0 (GradExpn A B0 (Suc n)) 
          col_transitivity_1 by blast
      moreover
      have "C' B' Perp A C0" 
        using Perp_perm A C0 Perp B' C' by blast
      hence "C' B' Perp A C'"
        using A  C' Col A C0 C' perp_col1 by blast
      hence "B' C' Perp C' A" 
        using Perp_cases by blast
      hence "Per B' C' A" 
        using C' B' Perp A C' perp_per_1 by blast
      moreover
      have "?B Midpoint A B'" 
        using Bet A (GradExpn A B0 (Suc n)) B' 
          Cong A (GradExpn A B0 (Suc n)) (GradExpn A B0 (Suc n)) B' 
          midpoint_def by blast
      moreover
      have "Col A C' C" 
        by (metis Col A C0 C' Col A C0 C ¬ Col A B0 C0 
            col_transitivity_1 not_col_distincts)
      ultimately
      show ?thesis using t22_23 
        using A  C' GradExpn A B0 (Suc n) Midpoint C D 
          Per A C (GradExpn A B0 (Suc n)) assms(1) l8_2 by blast
    qed
    ultimately
    have "B0 E' Le B' C'" 
      using le_transitivity by blast
  }
  thus ?thesis 
    by blast
qed

lemma t22_24_aux_n:
  assumes "¬ hypothesis_of_obtuse_saccheri_quadrilaterals"
  shows " A B0 C0 B C E.
  (¬ Col A B0 C0  A C0 Perp B0 C0  
  (B = GradExpn A B0 (Suc n))  (E = GradExpn B0 C0 (Suc n))  
  A C0 Perp B C  Col A C0 C) 
  B0 E Le B C" 
proof (induction n)
  case 0
  thus ?case 
    using t22_24_aux_0 by fastforce
next
  case (Suc n)
  thus ?case 
    using assms t22_24_aux_suc by presburger
qed

(** For every n, 2^n times B0C0 is lower than or equal to BnCn *)
(** B0 is introduced twice for the induction tactic to work properly *)
lemma t22_24_aux:
  assumes "¬ hypothesis_of_obtuse_saccheri_quadrilaterals"
  shows " A B0 B00 C0 B C E.
  ¬ Col A B0 C0  A C0 Perp B0 C0  B0 = B00 
  GradExp2 A B0 B B00 C0 E  A C0 Perp B C  Col A C0 C 
  B0 E Le B C"
proof -
  {
    fix A B0 B00 C0 B C E
    assume " ¬ Col A B0 C0" and
      "A C0 Perp B0 C0" and
      "B0 = B00" and
      "GradExp2 A B0 B B00 C0 E" and
      "A C0 Perp B C" and
      "Col A C0 C"
    obtain n where "n  0  (B = GradExpn A B0 n)  (E = GradExpn B00 C0 n)" 
      using GradExp2_def GradExp2 A B0 B B00 C0 E by presburger
    hence "n  0" 
      by blast
    then obtain m where "n = Suc m" 
      using not0_implies_Suc by blast
    have "B = GradExpn A B0 (Suc m)" 
      by (simp add: n = Suc m n  0  B = GradExpn A B0 n  E = GradExpn B00 C0 n)
    moreover
    have "E = GradExpn B0 C0 (Suc m)" 
      by (simp add:  B0 = B00 n = Suc m 
          n  0  B = GradExpn A B0 n  E = GradExpn B00 C0 n)
    ultimately
    have "B0 E Le B C" 
      using A C0 Perp B C A C0 Perp B0 C0 Col A C0 C 
        ¬ Col A B0 C0 assms t22_24_aux_n by blast
  }
  thus ?thesis by blast
qed

lemma t22_24_aux1_0:
  shows" A B0 C0 E. 
  (¬ Col A B0 C0  A C0 Perp B0 C0 
  E = GradExpn B0 C0 (Suc 0)) 
  ( B C. GradExp2 A B0 B B0 C0 E  A C0 Perp B C  Col A C0 C)"
proof -
  {
    fix A B0 C0 E
    assume "¬ Col A B0 C0" and
      "A C0 Perp B0 C0" and
      "E = GradExpn B0 C0 (Suc 0)"
    have "B0  C0" 
      using ¬ Col A B0 C0 not_col_distincts by blast
    hence "E = C0" 
      using E = GradExpn B0 C0 (Suc 0) by simp
    have "A  B0" 
      using ¬ Col A B0 C0 col_trivial_1 by blast
    hence "B0 = GradExpn A B0 (Suc 0)" 
      by simp
    have "GradExp2 A B0 B0 B0 C0 E" 
      using E = C0 gradexp2_init by auto
    moreover
    have "Col A C0 C0" 
      by (simp add: col_trivial_2)
    ultimately
    have " B C. GradExp2 A B0 B B0 C0 E  A C0 Perp B C  Col A C0 C" 
      using A C0 Perp B0 C0 by blast
  }
  thus ?thesis 
    by blast
qed

lemma t22_24_aux1_suc:
  assumes " A B0 C0 E. 
  (¬ Col A B0 C0  A C0 Perp B0 C0 
  E = GradExpn B0 C0 (Suc(n))) 
  ( B C. GradExp2 A B0 B B0 C0 E  A C0 Perp B C  Col A C0 C)"
  shows " A B0 C0 E. 
  (¬ Col A B0 C0  A C0 Perp B0 C0  
  E = GradExpn B0 C0 (Suc(Suc(n)))) 
  ( B C. GradExp2 A B0 B B0 C0 E  A C0 Perp B C  Col A C0 C)" 
proof -
  {
    fix A B0 C0 E'
    assume "¬ Col A B0 C0" and
      "A C0 Perp B0 C0" and
      "E' = GradExpn B0 C0 (Suc(Suc(n)))"
    let ?E = "GradExpn B0 C0 (Suc n)"
    have "Bet B0 ?E E'" 
      using GradExpn_3 E' = GradExpn B0 C0 (Suc (Suc n)) by blast
    have "Cong B0 ?E ?E E'" 
      using GradExpn_4 E' = GradExpn B0 C0 (Suc (Suc n)) by blast
    obtain B C where "GradExp2 A B0 B B0 C0 ?E  A C0 Perp B C  Col A C0 C" 
      using A C0 Perp B0 C0 ¬ Col A B0 C0 assms by presburger
    have "GradExp2 A B0 B B0 C0 ?E" 
      using GradExp2 A B0 B B0 C0 ?E  A C0 Perp B C  Col A C0 C by blast
    hence "GradExp A B0 B" 
      by (simp add: gradexp2__gradexp123)
    hence "Grad A B0 B" 
      by (simp add: gradexp__grad)
    hence "Bet A B0 B" 
      by (simp add: grad__bet)
    have "A C0 Perp B C" 
      using GradExp2 A B0 B B0 C0 ?E  A C0 Perp B C  Col A C0 C by blast
    have "Col A C0 C" 
      using GradExp2 A B0 B B0 C0 ?E  A C0 Perp B C  Col A C0 C by blast
    obtain B' where "Bet A B B'  Cong B B' A B" 
      using segment_construction by presburger
    have "Bet A B B'" 
      using Bet A B B'  Cong B B' A B by blast
    have "Cong B B' A B" 
      using Bet A B B'  Cong B B' A B by blast
    have "¬ Col A C0 B'" 
      by (metis Bet A B0 B Bet A B B' ¬ Col A B0 C0 
          bet_col bet_neq12__neq col_trivial_3 colx not_col_permutation_5)
    then obtain C' where "Col A C0 C'  A C0 Perp B' C'" 
      using l8_18_existence by blast
    moreover
    have "GradExp2 A B0 B' B0 C0 E'" 
    proof -
      have "GradExp2 A B0 B B0 C0 ?E" 
        using GradExp2 A B0 B B0 C0 (GradExpn B0 C0 (Suc n)) by blast
      moreover
      have "Bet A B B'" 
        by (simp add: Bet A B B'  Cong B B' A B)
      moreover
      have "Cong A B B B'" 
        using Cong_cases Cong B B' A B by blast
      moreover
      have "Bet B0 ?E E'" 
        using Bet B0 (GradExpn B0 C0 (Suc n)) E' by auto
      moreover
      have "Cong B0 ?E ?E E'" 
        using Cong B0 (GradExpn B0 C0 (Suc n)) (GradExpn B0 C0 (Suc n)) E' 
        by auto
      ultimately
      show ?thesis 
        using GradExp2_stab by blast
    qed
    ultimately
    have " B C. GradExp2 A B0 B B0 C0 E'  A C0 Perp B C  Col A C0 C" 
      by blast
  }
  thus ?thesis 
    by blast
qed

lemma t22_24_aux1_n:
  shows " A B0 C0 E. 
  (¬ Col A B0 C0  A C0 Perp B0 C0 
  E = GradExpn B0 C0 (Suc(n))) 
  ( B C. GradExp2 A B0 B B0 C0 E  A C0 Perp B C  Col A C0 C)"
proof (induction n)
  case 0
  thus ?case using t22_24_aux1_0 by blast
next
  case (Suc n)
  thus ?case using t22_24_aux1_suc by presburger
qed

(** For every n, it is possible to get Bn and Cn *)
lemma t22_24_aux1:
  assumes "¬ Col A B0 C0" and
    "A C0 Perp B0 C0" and
    "GradExp B0 C0 E"
  shows " B C. GradExp2 A B0 B B0 C0 E  A C0 Perp B C  Col A C0 C" 
proof -
  obtain n where "n  0  E = GradExpn B0 C0 n" 
    using GradExp_def assms(3) by presburger
  hence "n  0"
    by simp
  then obtain m where "n = Suc m" 
    using not0_implies_Suc by presburger
  thus ?thesis 
    using t22_24_aux1_n assms(1) assms(2) n  0  E = GradExpn B0 C0 n by blast
qed

lemma t22_24:
  assumes "archimedes_axiom"
  shows "aristotle_s_axiom"
proof -
  {
    fix P Q D A B0
    assume "¬ Col D A B0" and
      "Acute D A B0"
    obtain C0 where "Col A D C0" and "A D Perp B0 C0" 
      using Col_cases ¬ Col D A B0 l8_18_existence by blast
    hence "A  C0" 
      using Acute D A B0 acute_col_perp__out acute_sym 
        col_trivial_3 l6_3_1 by blast
    {
      assume "Col A B0 C0"
      hence "Col C0 A B0" 
        using Col_cases by blast
      hence "Col D A B0" 
        by (metis A D Perp B0 C0 Acute D A B0 
            acute_col_perp__out acute_sym col_trivial_3 l6_3_1 
            not_col_permutation_1 perp_col1)
      hence False 
        using ¬ Col D A B0 by blast
    }
    hence "¬ Col A B0 C0" 
      by blast
    have "A C0 Perp B0 C0" 
      using Col A D C0 A  C0 A D Perp B0 C0 perp_col by blast
    have " X Y. A Out D X  A Out B0 Y  Per A X Y  P Q Lt X Y" 
    proof (cases "P = Q")
      case True
      have "Acute B0 A D" 
        by (simp add: Acute D A B0 acute_sym)
      hence "A Out D C0" 
        using Col A D C0 A D Perp B0 C0 
          acute_col_perp__out l6_6 by blast
      moreover
      have "A Out B0 B0" 
        by (metis ¬ Col D A B0 col_trivial_2 out_trivial)
      moreover
      have "Per A C0 B0" 
        by (simp add: A C0 Perp B0 C0 perp_comm perp_per_2)
      moreover
      have "¬ Cong P P C0 B0" 
        using A C0 Perp B0 C0 cong_reverse_identity 
          perp_not_eq_2 by blast
      hence "P P Lt C0 B0" 
        by (metis cong_trivial_identity lt1123)
      ultimately
      show ?thesis
        using True by blast
    next
      case False
      obtain Q' where "Bet P Q Q'" and "Cong Q Q' P Q" 
        using segment_construction by blast
      have "B0  C0" 
        using ¬ Col A B0 C0 col_trivial_2 by blast
      hence "Reach B0 C0 P Q'" 
        using archimedes_axiom_def assms by blast
      then obtain E where "GradExp B0 C0 E" and "P Q' Le B0 E" 
        using reach__ex_gradexp_le by blast
      obtain B C where "GradExp2 A B0 B B0 C0 E  A C0 Perp B C  Col A C0 C" 
        using A C0 Perp B0 C0 GradExp B0 C0 E ¬ Col A B0 C0 
          t22_24_aux1 by blast
      have "GradExp2 A B0 B B0 C0 E" 
        using GradExp2 A B0 B B0 C0 E  A C0 Perp B C  Col A C0 C by blast
      hence "Grad A B0 B" 
        using gradexp2__gradexp123 gradexp__grad by blast
      hence "Bet A B0 B" 
        by (simp add: grad__bet)
      have "A C0 Perp B C" 
        using GradExp2 A B0 B B0 C0 E  A C0 Perp B C  Col A C0 C by blast
      have "Col A C0 C" 
        using GradExp2 A B0 B B0 C0 E  A C0 Perp B C  Col A C0 C by blast
      have "A Out B0 B" 
        by (metis Bet A B0 B ¬ Col D A B0 bet_out col_trivial_2)
      have "Acute D A B" 
        by (metis Acute D A B0 Bet A B0 B ¬ Col D A B0 
            acute_col__out acute_out2__acute acute_trivial bet2__out 
            bet_neq21__neq l5_1 not_col_distincts)
      {
        assume "A = C"
        have "Per D A B" 
        proof -
          have "A C0 Perp B A" 
            using A = C A C0 Perp B C by auto
          moreover
          have "Col A C0 D" 
            using Col_cases Col A D C0 by blast
          ultimately
          show ?thesis 
            by (meson col_trivial_3 l8_16_1 l8_2)
        qed
        hence False 
          using Acute D A B acute_not_per by auto
      }
      hence "A  C" 
        by blast
      have "A Out D C" 
      proof -
        have "Acute B A D" 
          by (simp add: Acute D A B acute_sym)
        moreover
        have "Col A D C" 
          using A  C0 Col A C0 C Col A D C0 col_trivial_3 
            colx by blast
        moreover
        have "A D Perp B C" 
          by (metis A C0 Perp B C Col A D C0 ¬ Col D A B0 
              col_trivial_3 not_col_permutation_5 perp_col)
        ultimately
        show ?thesis 
          using acute_col_perp__out l6_6 by blast
      qed
      moreover
      have "A Out B0 B" 
        by (simp add: A Out B0 B)
      moreover
      have "Per A C B" 
        using A C0 Perp B C Col A C0 C col_trivial_3 
          l8_16_1 l8_2 by blast
      moreover
      have "P Q Lt C B" 
      proof -
        have "P Q Lt P Q'" 
          by (metis False bet__lt1213 cong_diff_3 
              Bet P Q Q' Cong Q Q' P Q)
        moreover
        have "B0 E Le B C" 
        proof -
          have "¬ hypothesis_of_obtuse_saccheri_quadrilaterals" 
            by (simp add: archi__obtuse_case_elimination assms)
          moreover
          have "¬ Col A B0 C0" 
            by (simp add: ¬ Col A B0 C0)
          moreover
          have "A C0 Perp B0 C0" 
            by (simp add: A C0 Perp B0 C0)
          moreover
          have "GradExp2 A B0 B B0 C0 E" 
            by (simp add: GradExp2 A B0 B B0 C0 E)
          moreover
          have "A C0 Perp B C" 
            by (simp add: A C0 Perp B C)
          moreover
          have "Col A C0 C" 
            using Col A C0 C by auto
          ultimately
          show ?thesis 
            using t22_24_aux by blast
        qed
        hence "B0 E Le C B" 
          by (simp add: le_right_comm)
        hence "P Q' Le C B" 
          using P Q' Le B0 E le_transitivity by blast
        ultimately
        show ?thesis 
          using le3456_lt__lt by blast
      qed
      ultimately
      show ?thesis 
        by blast
    qed
  }
  thus ?thesis
    by (simp add: aristotle_s_axiom_def)
qed

subsection "Equivalence Grad / GradI (inductive)"

lemma Grad1__GradI:
  shows "GradI A B (Gradn A B 1)"
proof -
  have "B = Gradn A B 1" 
    by force
  thus ?thesis 
    using gradi_init by auto
qed

lemma Gradn__GradI:
  shows "GradI A B (Gradn A B (Suc n))" 
proof (induction n)
  show "GradI A B (Gradn A B (Suc 0))" 
    by (simp add: gradi_init)
  {
    fix n
    assume "GradI A B (Gradn A B (Suc n))"
    have "GradI A B (Gradn A B (Suc (Suc n)))" 
    proof(rule GradI.induct [where ?A ="A" and ?B="B"])
      show "GradI A B (Gradn A B (Suc (Suc n)))" 
        using Bet_Gradn_Gradn_Suc Cong_Gradn_Gradn_Suc GradI.simps 
          GradI A B (Gradn A B (Suc n)) by blast
      show "GradI A B B" 
        by (simp add: gradi_init)
      show "C C'. GradI A B C  GradI A B C  Bet A C C'  Cong A B C C'  GradI A B C'" 
        using gradi_stab by blast
    qed
  }
  thus "n. GradI A B (Gradn A B (Suc n))  GradI A B (Gradn A B (Suc (Suc n)))" 
    by blast
qed

lemma Grad__GradI:
  assumes "Grad A B C"
  shows "GradI A B C" 
proof (cases "B = C")
  case True
  thus ?thesis 
    by (simp add: gradi_init)
next
  case False
  obtain n where "n  0" and "C = Gradn A B n" 
    using Grad_def assms by blast
  then obtain m where "n = Suc m" 
    using not0_implies_Suc by blast
  hence "C = Gradn A B (Suc m)"
    using C = Gradn A B n by blast
  thus ?thesis 
    using Gradn__GradI by blast
qed

lemma GradIAAB__not:
  assumes "GradI A A B" 
  shows "A = B"
proof (rule GradI.induct [OF assms])
  show "A = A" 
    by blast
  show "C C'. GradI A A C  A = C  Bet A C C'  Cong A A C C'  A = C'"
    using between_cong by presburger
qed

lemma GradI__Grad:
  assumes "GradI A B C"
  shows "Grad A B C" 
proof (cases "A = B")
  case True
  hence "A = C" 
    using GradIAAB__not assms by blast
  thus ?thesis 
    using True grad_equiv_coq_1 by blast
next
  case False
  show ?thesis 
  proof (rule GradI.induct [where ?A="A" and ?B="B"])
    show "GradI A B C" 
      by (simp add: assms)
    show "Grad A B B" 
      by (simp add: grad_equiv_coq_1)
    show "C C'. GradI A B C  Grad A B C  Bet A C C'  Cong A B C C'  Grad A B C'" 
      using grad_stab by blast
  qed
qed

theorem Grad_GradI:
  shows "Grad A B C  GradI A B C" 
  using GradI__Grad Grad__GradI by blast

subsection "GradA"

lemma grada_distincts:
  assumes "GradA A B C D E F" 
  shows "A  B  C  B  D  E  F  E"
proof(induction rule: GradA.induct [OF assms(1)])
  show "D E F. A B C CongA D E F  A  B  C  B  D  E  F  E" 
    by (simp add: CongA_def)
  show "D E F G H I.
       GradA A B C D E F 
       A  B  C  B  D  E  F  E 
       SAMS D E F A B C 
       D E F A B C SumA G H I  A  B  C  B  G  H  I  H" 
    using suma_distincts by blast
qed

lemma grada_ABC:
  assumes "A  B" and
    "B  C"
  shows "GradA A B C A B C" 
proof -
  have "A B C CongA A B C" 
    using assms(1) assms(2) conga_refl by force
  thus ?thesis 
    using grada_init 
    by force
qed

lemma gradaexp_distincts:
  assumes "GradAExp A B C D E F" 
  shows "A  B  C  B  D  E  F  E"
proof(induction rule: GradAExp.induct [OF assms(1)])
  show "D E F. A B C CongA D E F  A  B  C  B  D  E  F  E" 
    by (simp add: CongA_def)
  show "D E F G H I.
       GradAExp A B C D E F 
       A  B  C  B  D  E  F  E 
       SAMS D E F D E F  D E F D E F SumA G H I  A  B  C  B  G  H  I  H"
    using suma_distincts by blast
qed

lemma gradaexp_ABC:
  assumes "A  B" and
    "B  C"
  shows "GradAExp A B C A B C" 
proof -
  have "A B C CongA A B C" 
    using assms(1) assms(2) conga_refl by force
  thus ?thesis 
    using gradaexp_init 
    by force
qed

lemma conga2_grada__grada_aux1: 
  assumes "GradA A B C D E F" and
    "A B C CongA A' B' C'" 
  shows "GradA A' B' C' D E F" 
proof (induction rule: GradA.cases [OF assms(1)])
  { 
    assume "A B C CongA D E F"
    hence "GradA A' B' C' D E F" 
      using grada_init 
      by (meson assms(2) not_conga not_conga_sym)
  }
  thus "Da Ea Fa.
       D = Da 
       E = Ea  F = Fa  A B C CongA Da Ea Fa  GradA A' B' C' D E F" 
    by blast
  {
    fix Da Ea Fa
    assume 1: "GradA A B C Da Ea Fa" and
      2: "SAMS Da Ea Fa A B C" and
      3: "Da Ea Fa A B C SumA D E F" 
    have "GradA A' B' C' D E F" 
    proof (rule grada_stab)
      let ?D = "Da"
      let ?E = "Ea"
      let ?F = "Fa"
      show "GradA A' B' C' ?D ?E ?F" 
      proof (rule GradA.inducts)
        show "GradA A B C Da Ea Fa" 
          by (simp add: "1")
        show "D E F. A B C CongA D E F  GradA A' B' C' D E F" 
          by (meson conga_trans not_conga_sym assms(2) grada_init)
        {
          fix D E F G H I
          assume "GradA A B C D E F" and
            A:  "GradA A' B' C' D E F" and
            B:  "SAMS D E F A B C" and
            C:  "D E F A B C SumA G H I"
          have "GradA A' B' C' G H I"  
          proof (rule grada_stab)
            show "GradA A' B' C' D E F" 
              by (simp add: A)
            show "SAMS D E F A' B' C'" using B 
              by (meson C conga2_sams__sams assms(2) sams2_suma2__conga123)
            show "D E F A' B' C' SumA G H I" using C 
              by (meson B conga3_suma__suma sams2_suma2__conga123 suma2__conga assms(2))
          qed
        }
        thus "D E F G H I.
       GradA A B C D E F 
       GradA A' B' C' D E F 
       SAMS D E F A B C 
       D E F A B C SumA G H I  GradA A' B' C' G H I" 
          by blast
      qed
      show "SAMS ?D ?E ?F A' B' C'" using 2 
        by (meson "3" sams2_suma2__conga123 assms(2) conga2_sams__sams)
      show "?D ?E ?F A' B' C' SumA D E F" using 3 
        by (meson "2" sams2_suma2__conga123 assms(2) conga3_suma__suma suma2__conga)
    qed
  }
  thus "Da Ea Fa G H I.
       D = G 
       E = H 
       F = I 
       GradA A B C Da Ea Fa 
       SAMS Da Ea Fa A B C 
       Da Ea Fa A B C SumA G H I  GradA A' B' C' D E F " 
    by blast
qed

lemma conga2_grada__grada_aux2: 
  assumes "GradA A B C D E F" and
    "D E F CongA D' E' F'" 
  shows "GradA A B C D' E' F'" 
proof (induction rule: GradA.cases [OF assms(1)])
  {
    assume "A B C CongA D E F"
    hence "A B C CongA D' E' F'" 
      by (metis conga_trans assms(2))
    hence "GradA A B C D' E' F'" 
      by (simp add: grada_init)
  }
  thus "Da Ea Fa.
       D = Da 
       E = Ea  F = Fa  A B C CongA Da Ea Fa  GradA A B C D' E' F'" 
    by blast
  {
    fix Da Ea Fa
    assume 1: "GradA A B C Da Ea Fa" and
      2: "SAMS Da Ea Fa A B C" and
      3: "Da Ea Fa A B C SumA D E F"
    have  "GradA A B C D' E' F'" 
    proof (rule grada_stab)
      show "GradA A B C Da Ea Fa" 
        by (simp add: "1")
      show "SAMS Da Ea Fa A B C" 
        by (simp add: "2")
      show "Da Ea Fa A B C SumA D' E' F'" 
        by (meson "2" "3" sams2_suma2__conga123 assms(2) 
            conga3_suma__suma sams2_suma2__conga456)
    qed
  }
  thus "Da Ea Fa G H I.
       D = G 
       E = H 
       F = I 
       GradA A B C Da Ea Fa 
       SAMS Da Ea Fa A B C 
       Da Ea Fa A B C SumA G H I  GradA A B C D' E' F'" 
    by blast
qed

lemma conga2_grada__grada: 
  assumes "GradA A B C D E F" and
    "A B C CongA A' B' C'" and
    "D E F CongA D' E' F'"
  shows "GradA A' B' C' D' E' F'" 
  using assms(1) assms(2) assms(3) 
    conga2_grada__grada_aux1 conga2_grada__grada_aux2 by blast

lemma grada__lea:
  assumes "GradA A B C D E F"
  shows "A B C LeA D E F" 
proof (induction rule: GradA.cases [OF assms(1)])
  case (1 D E F)
  then show ?case 
    by (simp add: conga__lea)
next
  case (2 D E F G H I)
  thus ?case 
    by (metis sams_suma__lea456789)
qed

lemma grada_out__out:
  assumes "E Out D F" and 
    "GradA A B C D E F" 
  shows "B Out A C" 
proof (induction rule: GradA.cases [OF assms(2)])
  case (1 D E F)
  then show ?case 
    by (metis not_conga_sym assms(1) l11_21_a)
next
  case (2 D E F G H I)
  then show ?case 
    by (metis sams_suma__lea456789 assms(1) out_lea__out)
qed

lemma grada2_sams_suma__grada_aux:
  shows " A B C D E F G H I K L M. 
              GradA A B C D E F  GradA A B C G H I  
              SAMS D E F G H I  D E F G H I SumA K L M  GradA A B C K L M"
proof -
  {
    fix A B C D E F G' H' I' K L M
    assume K1: "GradA A B C D E F" and
      K2: "GradA A B C G' H' I'" 
    have " K L M. (SAMS D E F G' H' I'  D E F G' H' I' SumA K L M)  GradA A B C K L M" 
    proof(induction rule: GradA.induct [OF K2])
      {
        fix Da Ea Fa
        assume "A B C CongA Da Ea Fa"
        {
          fix K L M
          assume "SAMS D E F Da Ea Fa" and "D E F Da Ea Fa SumA K L M"
          have "SAMS D E F A B C" 
            by (meson conga2_sams__sams not_conga_sym sams2_suma2__conga123 
                A B C CongA Da Ea Fa D E F Da Ea Fa SumA K L M SAMS D E F Da Ea Fa)
          moreover have "D E F A B C SumA K L M" 
            by (meson conga3_suma__suma not_conga_sym sams2_suma2__conga123 
                A B C CongA Da Ea Fa D E F Da Ea Fa SumA K L M SAMS D E F Da Ea Fa 
                suma2__conga)
          ultimately have "GradA A B C K L M" 
            using K1 grada_stab by blast
        }
        hence "K L M. (SAMS D E F Da Ea Fa  D E F Da Ea Fa SumA K L M)  GradA A B C K L M" 
          by blast
      }
      thus "Da Ea Fa.
       A B C CongA Da Ea Fa 
       K L M. (SAMS D E F Da Ea Fa  D E F Da Ea Fa SumA K L M)  GradA A B C K L M" 
        by blast
      {
        fix G H I 
        {
          fix Da Ea Fa
          assume "GradA A B C Da Ea Fa" and
            "SAMS Da Ea Fa A B C" and
            "Da Ea Fa A B C SumA G H I" and
            P1: "K L M. SAMS D E F Da Ea Fa  D E F Da Ea Fa SumA K L M  GradA A B C K L M"
          {
            fix K0 L0 M0
            assume "SAMS D E F G H I" and
              "D E F G H I SumA K0 L0 M0"
            have "Da  Ea" 
              using SAMS Da Ea Fa A B C sams_distincts by auto
            have "Fa  Ea" 
              using GradA A B C Da Ea Fa grada_distincts by blast
            have "D  E" 
              using SAMS D E F G H I sams_distincts by blast
            have "E  F" 
              using SAMS D E F G H I sams_distincts by blast
            obtain K L M where "D E F Da Ea Fa SumA K L M" 
              using ex_suma D  E Da  Ea E  F Fa  Ea by presburger
            have "SAMS D E F Da Ea Fa" 
            proof (rule sams_lea2__sams [where ?A'="D" and ?B'="E" and ?C'="F" and
                  ?D'="G" and ?E'="H" and ?F'="I"]) 
              show "SAMS D E F G H I" 
                using SAMS D E F G H I by blast
              show "D E F LeA D E F" 
                using D  E E  F lea_refl by force
              show "Da Ea Fa LeA G H I"
              proof (rule sams_suma__lea123789 [where ?D="A" and ?E="B" and ?F="C"])
                show "Da Ea Fa A B C SumA G H I" 
                  using Da Ea Fa A B C SumA G H I by blast
                show "SAMS Da Ea Fa A B C"
                  by (simp add: SAMS Da Ea Fa A B C)
              qed
            qed
            have "GradA A B C K0 L0 M0" 
            proof (rule grada_stab [where ?D = "K" and ?E = "L" and ?F = "M"])
              show "GradA A B C K L M" 
                using P1 D E F Da Ea Fa SumA K L M SAMS D E F Da Ea Fa by blast
              show "SAMS K L M A B C" 
                using sams_assoc_2 [where ?A="D" and ?B="E" and ?C="F" and
                    ?D="Da" and ?E="Ea" and ?F="Fa" and ?D'="G" and ?E'="H" and ?F'="I"] 
                using D E F Da Ea Fa SumA K L M Da Ea Fa A B C SumA G H I 
                  SAMS D E F Da Ea Fa SAMS D E F G H I SAMS Da Ea Fa A B C by blast
              show "K L M A B C SumA K0 L0 M0" 
                by (meson suma_assoc_2 D E F Da Ea Fa SumA K L M 
                    D E F G H I SumA K0 L0 M0 Da Ea Fa A B C SumA G H I SAMS D E F Da Ea Fa 
                    SAMS Da Ea Fa A B C)
            qed
          }
          hence "K L M. SAMS D E F G H I  D E F G H I SumA K L M  GradA A B C K L M" 
            by blast
        }
        hence " Da Ea Fa. GradA A B C Da Ea Fa  SAMS Da Ea Fa A B C 
                            Da Ea Fa A B C SumA G H I 
     (K L M. SAMS D E F Da Ea Fa  D E F Da Ea Fa SumA K L M  GradA A B C K L M) 
     (K L M. SAMS D E F G H I  D E F G H I SumA K L M  GradA A B C K L M)" 
          by blast
      }
     thus "Da Ea Fa G H I.
       GradA A B C Da Ea Fa 
       (K L M. SAMS D E F Da Ea Fa  D E F Da Ea Fa SumA K L M  GradA A B C K L M) 
       SAMS Da Ea Fa A B C 
       Da Ea Fa A B C SumA G H I 
       (K L M. SAMS D E F G H I  D E F G H I SumA K L M  GradA A B C K L M)" 
        by blast
    qed
  }
  thus ?thesis by blast
qed

lemma grada2_sams_suma__grada:
  assumes "GradA A B C D E F" and
    "GradA A B C G H I" and
    "SAMS D E F G H I" and 
    "D E F G H I SumA K L M"
  shows "GradA A B C K L M" 
  using assms(1) assms(2) assms(3) assms(4) grada2_sams_suma__grada_aux by blast

lemma gradaexp__grada:
  assumes "GradAExp A B C D E F"
  shows "GradA A B C D E F" 
proof (rule GradAExp.induct [OF assms])
  show "D E F. A B C CongA D E F  GradA A B C D E F" 
    by (simp add: grada_init)
  show " D E F G H I.
       GradAExp A B C D E F 
       GradA A B C D E F  SAMS D E F D E F  D E F D E F SumA G H I  GradA A B C G H I" 
    using grada2_sams_suma__grada_aux by blast
qed

lemma acute_archi_aux:
  assumes "Per PO A B" and
    "PO  A" and
    "B  A" and
    "C  D" and
    "D  E" and
    "Bet A C D" and
    "Bet C D E" and
    "Bet D E B" and
    "C PO D CongA D PO E"
  shows "C D Lt D E"
proof -
  have "D  A" 
    using assms(4) assms(6) between_identity by blast
  have "C  PO" 
    using assms(9) conga_diff1 by auto
  have "D  PO" 
    using assms(9) conga_diff45 by blast
  have "¬ Col PO A B" 
    by (metis assms(1) assms(2) assms(3) l8_8 not_col_permutation_2 per_col)
  hence "¬ Col PO A D" 
    by (metis Col_def D  A assms(4) assms(5) assms(6) assms(7) assms(8) l6_16_1)
  have "¬ Col PO D E" 
    by (metis Col_def ¬ Col PO A D assms(4) assms(5) assms(6) assms(7) l6_16_1)
  then obtain P where "A D PO CongA PO D P" and "PO D OS P E" 
    using ¬ Col PO A D angle_construction_1 not_col_permutation_1 by blast
  have "Acute A D PO" 
    by (metis D  A assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) 
        assms(8) bet_col1 between_trivial l11_43 outer_transitivity_between per_col)
  hence "Acute A D PO  A D PO LtA E D PO" 
    by (metis acute_chara_1 assms(4) assms(5) assms(6) assms(7) outer_transitivity_between2)
  have "A D PO LeA E D PO" 
    by (metis lta__lea outer_transitivity_between2 Acute A D PO acute_chara_1 
        assms(4) assms(5) assms(6) assms(7))
  hence "PO D P LeA PO D E" 
    by (meson lea_right_comm lea_trans A D PO CongA PO D P conga__lea456123)
  hence "P InAngle PO D E" 
    using PO D OS P E lea_in_angle one_side_symmetry by presburger
  have "P  D" 
    using PO D OS P E os_distincts by blast
  obtain F where "Bet PO F E" and "D Out F P" 
    using InAngle_def P InAngle PO D E ¬ Col PO D E bet_col by auto
  have "A D PO CongA PO D F" 
    by (metis out2__conga A D PO CongA PO D P D Out F P D  PO 
        bet_out_1 conga_trans not_bet_distincts)
  have "D Out A C" 
    by (simp add: assms(4) assms(6) bet_out_1 l6_6)
  have "D Out PO PO" 
    using D  PO out_trivial by auto
  have "PO D C CongA A D PO"
    by (simp add: out2__conga D Out A C D Out PO PO conga_left_comm)
  have "¬ Col PO D F" 
    using D Out F P D Out PO PO PO D OS P E col_out2_col col_trivial_3 l9_19 by blast
  have "PO  F" 
    using ¬ Col PO D F col_trivial_3 by auto
  have "¬ Col PO D C" 
    using ¬ Col PO A D assms(4) assms(6) bet_col col2__eq col_permutation_5 by blast
  have "PO Out D D" 
    by (simp add: D  PO out_trivial)
  have "PO Out C C" 
    by (simp add: C  PO out_trivial)
  have "PO Out F E" 
    using Bet PO F E PO  F bet_out by auto
  hence "D PO C CongA D PO F" 
    using l11_10 PO Out C C PO Out D D assms(9) conga_left_comm by blast
  have "PO D C CongA PO D F" 
    by (meson conga_trans A D PO CongA PO D F PO D C CongA A D PO)
  have "Cong PO D PO D" 
    by (simp add: cong_reflexivity)
  have "Cong PO C PO F  Cong D C D F  PO C D CongA PO F D" 
    using l11_50_1 Cong PO D PO D D PO C CongA D PO F PO D C CongA PO D F
      ¬ Col PO D C by blast
  hence "Cong D F C D" 
    using not_cong_4312 by blast
  moreover
  {
    assume "Col E D F"
    {
      assume "E = F"
      have "D  C" 
        using assms(4) by auto
      moreover have "Per PO D C" 
        using E = F PO D C CongA PO D F assms(7) l11_18_2 by auto
      moreover have "Col D C A" 
        using D Out A C not_col_permutation_5 out_col by blast
      ultimately have "Per A D PO" 
        by (meson l11_17 PO D C CongA A D PO)
      hence False 
        using acute_not_per Acute A D PO by blast
    }
    moreover
    {
      assume "E  F" 
      hence False
        by (metis col_permutation_1 Bet PO F E Col E D F ¬ Col PO D F bet_col col2__eq)
    }
    ultimately have False 
      by blast
  }
  hence "¬ Col E D F" 
    by blast
  have "E  F" 
    using ¬ Col E D F not_col_distincts by blast
  have "D E F LtA D F E" 
  proof (rule lta_trans [where ?A1.0 = "F" and ?B1.0 = "D" and ?C1.0 = "PO"])
    have "D PO E LtA PO D C  D E PO LtA PO D C" 
      by (metis l11_41 not_col_permutation_1 one_side_not_col124 
          thesis. (P. A D PO CongA PO D P; PO D OS P E  thesis)  thesis 
          assms(4) assms(7) between_symmetry)
    show "D E F LtA F D PO" 
    proof -
      have "D E PO CongA D E F" 
        by (metis bet_out_1 out2__conga Bet PO F E E  F assms(5) out_trivial)
      moreover have "PO D C CongA F D PO" 
        by (simp add: PO D C CongA PO D F conga_right_comm)
      ultimately show ?thesis
        by (simp add: conga_preserves_lta D PO E LtA PO D C  D E PO LtA PO D C)
    qed
    show "F D PO LtA D F E"     
      by (metis bet_col col_lta__bet l11_41_aux not_col_permutation_2 
          Bet PO F E D PO E LtA PO D C  D E PO LtA PO D C E  F PO D C CongA A D PO 
          PO D C CongA PO D F 
          thesis. (P. A D PO CongA PO D P; PO D OS P E  thesis)  thesis 
          bet_conga__bet col123__nos ncol_conga_ncol)
  qed
  hence "D F Lt D E" 
    using l11_44_2_b by blast
  ultimately show ?thesis 
    using cong2_lt__lt 
    using cong_reflexivity by blast
qed

lemma acute_archi_aux1:
  assumes "Per PO A0 B" and
    "B  A0" and
    "Bet A0 A1 B" and
    "GradA A0 PO A1 P Q R" and
    "A0  A1" 
  shows "A0 PO B LeA P Q R  ( A. Bet A0 A1 A  Bet A0 A B  P Q R CongA A0 PO A)" 
proof -
  have "A0  PO" 
    using assms(4) grada_distincts by auto
  have "A1  PO"  
    using assms(4) grada_distincts by auto
  have "P  Q" 
    using assms(4) grada_distincts by auto
  have "R  Q" 
    using assms(4) grada_distincts by auto
  have "¬ Col PO A0 B" 
    by (metis per_not_col A0  PO assms(1) assms(2))
  have "PO  B" 
    using assms(1) assms(2) per_distinct_1 by blast
  {
    assume "P Q R LeA A0 PO B"
    {
      assume "Col P Q R" 
      {
        assume "Q Out P R"
        hence "PO Out A0 A1"
          using grada_out__out [where ?D = "P" and ?E = "Q" and ?F = "R"] assms(4) by blast
        hence False 
          using ¬ Col PO A0 B assms(3) assms(4) assms(5) bet_col col_trivial_2 
            colx out_col by blast
      }
      hence "¬ Q Out P R" 
        by blast
      hence "Bet P Q R" 
        using Col P Q R not_out_bet by blast
      hence "Bet A0 PO B" 
        using P Q R LeA A0 PO B bet_lea__bet [where ?A = "P" and ?B = "Q" and ?C = "R"]
        by blast
      hence False 
        using ¬ Col PO A0 B bet_col not_col_permutation_4 by blast
    }
    hence "¬ Col P Q R" 
      by blast
    then obtain C where "P Q R CongA A0 PO C" and "A0 PO OS C B" 
      by (metis NCol_cases ¬ Col PO A0 B angle_construction_1)
    have "C InAngle A0 PO B" 
    proof (rule lea_in_angle)
      have "A0 PO B CongA A0 PO B" 
        using A0  PO PO  B conga_refl by auto
      thus "A0 PO C LeA A0 PO B" 
        using l11_30 [where ?A="P" and ?B="Q" and ?C="R" and ?D="A0" and ?E="PO" and ?F="B"]
          P Q R LeA A0 PO B P Q R CongA A0 PO C by blast
      show "A0 PO OS B C" 
        by (simp add: A0 PO OS C B one_side_symmetry)
    qed
    have "C  PO" 
      using A0 PO OS C B os_distincts by blast
    obtain A where "Bet A0 A B" and "A = PO  PO Out A C" 
      using InAngle_def C InAngle A0 PO B by blast
    hence "PO Out A C" 
      using Bet_cases Col_def ¬ Col PO A0 B by blast
    have "P Q R CongA A0 PO A" 
    proof (rule l11_10 [where ?A ="P" and ?C="R" and ?D="A0" and ?F="C"],
        insert P Q R CongA A0 PO C PO Out A C) 
      show "Q Out P P" 
        by (simp add: P  Q out_trivial)
      show "Q Out R R" 
        using R  Q out_trivial by auto
      show "PO Out A0 A0" 
        by (simp add: A0  PO out_trivial)
    qed
    have "Bet A0 A1 A" 
    proof (cases "A0 = A1")
      case True
      thus ?thesis
        using assms(5) by auto
    next
      case False
      hence "¬ Col A0 PO A" 
        by (metis Col_perm A0 PO OS C B PO Out A C l6_16_1 
            one_side_not_col123 out_col out_distinct)
      have "A1 InAngle A0 PO A" 
      proof (rule lea_in_angle)
        show "A0 PO A1 LeA A0 PO A"
        proof (rule l11_30 [where ?A="A0" and ?B="PO" and ?C="A1" and ?D="P" and ?E="Q" and ?F="R"])
          show "A0 PO A1 LeA P Q R" 
            by (simp add: assms(4) grada__lea)
          show "A0 PO A1 CongA A0 PO A1" 
            by (simp add: A0  PO A1  PO conga_refl)
          show "P Q R CongA A0 PO A"
            by (simp add: P Q R CongA A0 PO A)
        qed
        show "A0 PO OS A A1"
          using False Bet A0 A B ¬ Col A0 PO A assms(3) bet2__out 
            not_col_distincts out_one_side by presburger
      qed
      obtain X where "Bet A0 X A" and "X = PO  PO Out X A1" 
        using InAngle_def A1 InAngle A0 PO A by blast
      hence "PO Out X A1" 
        using ¬ Col A0 PO A bet_col by blast
      have "X = A1  Bet A0 A1 A" 
        using Bet A0 X A by blast
      moreover
      have "X  A1  Bet A0 A1 A" 
        by (meson Bet A0 A B Bet A0 X A PO Out X A1 ¬ Col A0 PO A 
            assms(3) bet_col bet_col1 col_permutation_2 colx out_col)
      ultimately
      show ?thesis 
        by blast
    qed
    moreover have "Bet A0 A B" 
      by (simp add: Bet A0 A B)
    ultimately have  "A0 PO B LeA P Q R  ( A. Bet A0 A1 A  Bet A0 A B  P Q R CongA A0 PO A)" 
      using P Q R CongA A0 PO A by blast
  }
  thus ?thesis 
    by (metis A0  PO P  Q PO  B R  Q lea_total)
qed


lemma acute_archi_aux2_1_a:
  assumes "Per PO A0 B" and
    "PO  A0" and
    "B  A0" and
    "Bet A0 A1 B" and
    "A0  A1" and "¬ Col PO A0 B" and "¬ Col A0 PO A1" and "PO  A1" and "PO  B"
  shows " P Q R. (GradA A0 PO A1 P Q R  (A0 PO B LeA P Q R 
                  ( A'. Bet A0 A1 A'  Bet A0 A' B  P Q R CongA A0 PO A'  A0 A1 Le A0 A' 
                  ( A. Bet A0 A A'  A0 PO A' CongA A0 PO A1  A0 A1 Le A A'))))" 
proof -
  let ?P = "A0"
  let ?Q = "PO"
  let ?R = "A1"
  have "(GradA A0 PO A1 ?P ?Q ?R  (A0 PO B LeA ?P ?Q ?R 
                  ( A'. Bet A0 A1 A'  Bet A0 A' B  ?P ?Q ?R CongA A0 PO A'  A0 A1 Le A0 A' 
                  ( A. Bet A0 A A'  A0 PO A' CongA A0 PO A1  A0 A1 Le A A'))))" 
    by (metis assms(2) assms(4) assms(8) bet__le1213 conga_refl grada_init not_bet_distincts)
  thus ?thesis by blast
qed

lemma acute_archi_aux2_1:
  assumes "Per PO A B" and
    "PO  A" and
    "B  A" and
    "Bet A B0 B" and
    "A  B0" and "¬ Col PO A B" and "¬ Col A PO B0" and "PO  B0" and "PO  B"
  shows " P Q R. (GradA A PO B0 P Q R  (A PO B LeA P Q R 
                  ( A'. Bet A B0 A'  Bet A A' B  P Q R CongA A PO A'  A B0 Le A A' 
                  ( A0. Bet A A0 A'  A0 PO A' CongA A PO B0  A B0 Le A0 A'))))" 
proof -
  let ?A0 = "A"
  let ?A1 = "B0"
  have " P Q R. (GradA ?A0 PO ?A1 P Q R  (?A0 PO B LeA P Q R 
                  ( A'. Bet ?A0 ?A1 A'  Bet ?A0 A' B  P Q R CongA ?A0 PO A'  
                         ?A0 ?A1 Le ?A0 A' 
                  ( A. Bet ?A0 A A'  ?A0 PO A' CongA ?A0 PO ?A1  ?A0 ?A1 Le A A'))))" 
    using acute_archi_aux2_1_a assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) 
      assms(8) assms(9) by blast
  thus ?thesis 
    by (metis not_bet_distincts)
qed

lemma acute_archi_aux2_2:
  assumes "Per PO A0 B" and
    "PO  A0" and
    "B  A0" and
    "Bet A0 A1 B" and
    "A0  A1" and
    "Grad A0 A1 C" and
    "Bet A0 C C'" and
    "Cong A0 A1 C C'" and
    "¬ Col PO A0 B" and
    "¬ Col A0 PO A1" and
    "PO  A1" and
    "PO  B" and
    "Per PO A0 B  PO  A0 
       B  A0  Bet A0 A1 B 
       A0  A1 
       ¬ Col PO A0 B 
       ¬ Col A0 PO A1 
       PO  A1  ( P Q R.
         (GradA A0 PO A1 P Q R 
         (A0 PO B LeA P Q R 
          ( A'.
             Bet A0 A1 A' 
             Bet A0 A' B 
             P Q R CongA A0 PO A' 
             A0 C Le A0 A'  ( A. ( Bet A0 A A'  A PO A' CongA A0 PO A1  A0 A1 Le A A'))))))"
  shows " P Q R. (GradA A0 PO A1 P Q R  
(A0 PO B LeA P Q R 
( A'. Bet A0 A1 A'  Bet A0 A' B  P Q R CongA A0 PO A'  A0 C' Le A0 A' 
                     ( A. Bet A0 A A'  A PO A' CongA A0 PO A1  A0 A1 Le A A'))))" 
proof -
  have " P Q R.
         (GradA A0 PO A1 P Q R  (A0 PO B LeA P Q R 
          ( A'. Bet A0 A1 A'  Bet A0 A' B  P Q R CongA A0 PO A' 
             A0 C Le A0 A'  ( A. ( Bet A0 A A'  A PO A' CongA A0 PO A1  A0 A1 Le A A')))))"
    using assms(13) assms(1) assms(10) assms(11) assms(2) assms(3) assms(4) 
      assms(5) assms(9) by blast
  then obtain P Q R where "GradA A0 PO A1 P Q R" and
    P2:  "A0 PO B LeA P Q R 
          ( A'. Bet A0 A1 A'  Bet A0 A' B  P Q R CongA A0 PO A'  A0 C Le A0 A'  
             ( A. ( Bet A0 A A'  A PO A' CongA A0 PO A1  A0 A1 Le A A')))"
    by blast
  {
    assume "A0 PO B LeA P Q R" 
    hence " P Q R. (GradA A0 PO A1 P Q R  (A0 PO B LeA P Q R 
               ( A'. Bet A0 A1 A'  Bet A0 A' B  P Q R CongA A0 PO A'  A0 C' Le A0 A' 
                     ( A. Bet A0 A A'  A PO A' CongA A0 PO A1  A0 A1 Le A A'))))" 
      using GradA A0 PO A1 P Q R by blast
  }
  moreover
  {
    assume "¬ A0 PO B LeA P Q R" 
    assume " A'.
             Bet A0 A1 A' 
             Bet A0 A' B 
             P Q R CongA A0 PO A' 
             A0 C Le A0 A'  ( A. ( Bet A0 A A'  A PO A' CongA A0 PO A1  A0 A1 Le A A'))"
    then obtain A' where "Bet A0 A1 A'" and
      "Bet A0 A' B" and
      "P Q R CongA A0 PO A'" and
      "A0 C Le A0 A'" and 
      P3: "( A. ( Bet A0 A A'  A PO A' CongA A0 PO A1  A0 A1 Le A A'))"
      by blast
    then obtain A where "Bet A0 A A'" and "A PO A' CongA A0 PO A1" and "A0 A1 Le A A'"
      by blast
    have "SAMS P Q R A0 PO A1" 
    proof (rule sams_lea2__sams [where ?A'="A0" and ?B'="PO" and ?C'="B" and 
          ?D'="A0" and ?E'="PO" and ?F'="B"])
      show "SAMS A0 PO B A0 PO B" 
        by (metis acute_chara_1 lea_right_comm lta__lea assms(1) assms(2) assms(3) l11_43 
            point_construction_different sams_chara)
      show "P Q R LeA A0 PO B"
      proof (rule l11_30 [where ?A="A0" and ?B="PO" and ?C="A'" and 
            ?D="A0" and ?E="PO" and ?F="B"])
        have "A' InAngle A0 PO B" 
          by (metis InAngle_def Out_def A PO A' CongA A0 PO A1 
              Bet A0 A' B assms(12) assms(2) between_trivial conga_diff2)
        thus "A0 PO A' LeA A0 PO B" 
          using inangle__lea by force
        show "A0 PO A' CongA P Q R" 
          using P Q R CongA A0 PO A' conga_sym_equiv by auto
        show "A0 PO B CongA A0 PO B" 
          using assms(12) assms(2) conga_refl by force
      qed
      have "A1 InAngle A0 PO B" 
        using InAngle_def assms(11) assms(12) assms(2) assms(4) out_trivial by auto
      thus "A0 PO A1 LeA A0 PO B" 
        by (simp add: inangle__lea)
    qed
    have "A0  A'" 
      using Bet A0 A1 A' assms(5) between_identity by blast
    have "A  A'" 
      using Le_def A0 A1 Le A A' assms(5) between_identity cong_identity_inv by blast
    have "PO  A" 
      using A PO A' CongA A0 PO A1 conga_diff1 by blast
    have "P  Q" 
      using P Q R CongA A0 PO A' conga_diff1 by blast
    have "PO  A'" 
      using P3 conga_diff2 by blast
    have "Q  R" 
      using P Q R CongA A0 PO A' conga_diff2 by blast
    then obtain P' Q' R' where "P Q R A0 PO A1 SumA P' Q' R'" 
      using ex_suma P  Q assms(11) assms(2) by fastforce
    have "GradA A0 PO A1 P' Q' R'" 
      using grada_stab [where ?D="P" and ?E="Q" and ?F="R"]
        GradA A0 PO A1 P Q R SAMS P Q R A0 PO A1 P Q R A0 PO A1 SumA P' Q' R' by blast
    moreover
    have "A0 PO B LeA P' Q' R'  ( A. Bet A0 A1 A  Bet A0 A B  P' Q' R' CongA A0 PO A)"
      using acute_archi_aux1 assms(1) assms(3) assms(4) assms(5) calculation by blast 
    moreover
    {
      assume "A0 PO B LeA P' Q' R'"
      hence "GradA A0 PO A1 P' Q' R'" 
        using calculation by auto
      hence " P Q R. (GradA A0 PO A1 P Q R  
(A0 PO B LeA P Q R 
( A'. Bet A0 A1 A'  Bet A0 A' B  P Q R CongA A0 PO A'  A0 C' Le A0 A' 
                     ( A. Bet A0 A A'  A PO A' CongA A0 PO A1  A0 A1 Le A A'))))" 
        using A0 PO B LeA P' Q' R' by blast
    }
    moreover
    {
      assume " A. Bet A0 A1 A  Bet A0 A B  P' Q' R' CongA A0 PO A"
      then obtain A'' where "Bet A0 A1 A''" and "Bet A0 A'' B" and "P' Q' R' CongA A0 PO A''" 
        by blast
      have "¬ Col A PO A'" 
        by (meson A  A' Bet A0 A A' Bet A0 A' B Bet A0 A1 A' assms(10) assms(4) 
            bet_col1 colx not_col_permutation_5)
      have "¬ Col A0 PO A'" 
        by (meson A0  A' Bet A0 A' B assms(10) assms(4) bet_col1 col_trivial_3 colx)
      {
        assume "Col A' PO A''" 
        have "A'  A''  False" 
          by (meson Bet A0 A1 A'' Bet A0 A1 A' Col A' PO A'' assms(10) 
              bet_col col_permutation_5 colx)
        moreover
        {
          assume "A' = A''"
          have "A0 PO A' A0 PO A1 SumA A0 PO A'" 
            using conga3_suma__suma [where ?A="P" and ?B="Q" and ?C="R" and
                ?D="A0" and ?E="PO" and ?F="A1" and ?G="P'" and ?H="Q'" and ?I="R'"]
            using A' = A'' P Q R A0 PO A1 SumA P' Q' R' P Q R CongA A0 PO A' 
              P' Q' R' CongA A0 PO A'' assms(11) assms(2) conga_refl by force
          have "¬ Col A0 PO A''" 
            using A' = A'' ¬ Col A0 PO A' by auto
          have "Bet A0 A' A''" 
            by (simp add: A' = A'' between_trivial)
          have "SAMS A0 PO A' A0 PO A1" 
            using conga2_sams__sams [where ?A="P" and ?B="Q" and ?C="R" and
                ?D="A0" and ?E="PO" and ?F="A1"] 
            using P Q R CongA A0 PO A' SAMS P Q R A0 PO A1 assms(11) assms(2) 
              conga_refl by presburger
          hence "Col A0 PO A1"
            using sams_suma__out546
            by (meson not_col_permutation_4 A0 PO A' A0 PO A1 SumA A0 PO A' out_col)
          hence False 
            using assms(10) by blast
        }
        ultimately have False 
          by blast
      }
      hence "¬ Col A' PO A''" 
        by blast
      have "¬ Col A0 PO A''" 
        by (metis Bet A0 A'' B Bet A0 A1 A'' assms(10) assms(4) bet_col1 
            between_identity colx not_col_distincts)
      have "Bet A0 A' A''" 
      proof (rule col_two_sides_bet [where ?B="PO"])
        show "Col A' A0 A''" 
          using Bet A0 A' B Bet A0 A'' B bet_col1 not_col_permutation_1 by blast
        have "A' InAngle A0 PO A''" 
        proof (rule lea_in_angle)
          show "A0 PO A' LeA A0 PO A''" 
            using l11_30 [where ?A="P" and ?B="Q" and ?C="R" and
                ?D="P'" and ?E="Q'" and ?F="R'"] 
            by (meson sams_suma__lea123789 P Q R A0 PO A1 SumA P' Q' R' 
                P Q R CongA A0 PO A' P' Q' R' CongA A0 PO A'' SAMS P Q R A0 PO A1)
          show "A0 PO OS A'' A'" 
            by (metis A0  A' Bet A0 A' B Bet A0 A'' B Bet A0 A1 A'' 
                ¬ Col A0 PO A' assms(5) bet2__out between_identity out_one_side)
        qed
        thus "A' PO TS A0 A''" 
          by (simp add: ¬ Col A' PO A'' ¬ Col A0 PO A' in_angle_two_sides 
              not_col_permutation_1 not_col_permutation_4)
      qed
      have "A PO A' CongA A' PO A''" 
      proof (rule conga_trans [where ?A'="A0" and ?B'="PO" and ?C'="A1"])
        have "PO  A''" 
          using ¬ Col A' PO A'' not_col_distincts by blast
        have "A'  A''" 
          using ¬ Col A' PO A'' not_col_distincts by blast
        have "¬ PO A' OS A0 A''" 
          using Bet A0 A' A'' col_trivial_3 one_side_chara by force
        show "A PO A' CongA A0 PO A1" 
          by (simp add: A PO A' CongA A0 PO A1)
        show "A0 PO A1 CongA A' PO A''" 
        proof (rule sams2_suma2__conga456 [where ?A="P" and ?B="Q" and ?C="R" 
              and ?G="P'" and ?H="Q'" and ?I="R'"])
          show "SAMS P Q R A0 PO A1" 
            using SAMS P Q R A0 PO A1 by auto
          show "SAMS P Q R A' PO A''" 
          proof (rule conga2_sams__sams [where ?A="A0" and ?B="PO" and ?C="A'" and 
                ?D="A'" and ?E="PO" and ?F="A''"])
            show "A0 PO A' CongA P Q R" 
              using P Q R CongA A0 PO A' conga_sym_equiv by blast
            show "A' PO A'' CongA A' PO A''" 
              using PO  A'' PO  A' conga_refl by auto
            show "SAMS A0 PO A' A' PO A''" 
              by (metis Col_cases bet_out bet_out_1 os2__sams out_one_side 
                  A'  A'' A0  A' Bet A0 A' A'' ¬ Col A0 PO A'')
          qed
          show "P Q R A0 PO A1 SumA P' Q' R'" 
            by (simp add: P Q R A0 PO A1 SumA P' Q' R')
          show "P Q R A' PO A'' SumA P' Q' R'"
          proof (rule conga3_suma__suma [where ?A="A0" and ?B="PO" and ?C="A'" and
                ?D="A'" and ?E="PO" and ?F="A''" and ?G="A0" and ?H="PO" and ?I="A''"])
            show "A0 PO A' A' PO A'' SumA A0 PO A''" 
            proof -
              have "A' PO A'' CongA A' PO A''"                 
                using PO  A'' PO  A' conga_refl by auto
              moreover have "¬ PO A' OS A0 A''"                   
                using ¬ PO A' OS A0 A'' by auto
              moreover have "Coplanar A0 PO A' A''"                   
                using Bet A0 A' A'' bet_col ncop__ncols by blast
              moreover have "A0 PO A'' CongA A0 PO A''"       
                using PO  A'' assms(2) conga_refl by auto
              ultimately show ?thesis
                using SumA_def by blast
            qed
            show "A0 PO A' CongA P Q R" 
              using P Q R CongA A0 PO A' conga_sym_equiv by auto
            show "A' PO A'' CongA A' PO A''" 
              using PO  A'' PO  A' conga_refl by auto
            show "A0 PO A'' CongA P' Q' R'"
              using P' Q' R' CongA A0 PO A'' conga_sym_equiv by auto
          qed
        qed
      qed
      have "A A' Lt A' A''" 
        using acute_archi_aux [where ?PO="PO" and ?A="A0" and ?B="B"] 
        by (metis ncol_conga_ncol not_col_distincts not_conga_sym 
            A PO A' CongA A' PO A'' A PO A' CongA A0 PO A1 Bet A0 A A'
            Bet A0 A' A'' Bet A0 A'' B assms(1) assms(10) assms(3) between_exchange3)
      hence "A A' Le A' A''" 
        using Lt_def by blast
      hence "A0 A1 Le A' A''" 
        by (meson le_transitivity A0 A1 Le A A')
      have "Bet A0 A1 A''" 
        by (simp add: Bet A0 A1 A'')
      moreover have "Bet A0 A'' B" 
        by (simp add: Bet A0 A'' B)
      moreover have "P' Q' R' CongA A0 PO A''" 
        by (simp add: P' Q' R' CongA A0 PO A'')
      moreover have "A0 C' Le A0 A''" 
        by (meson bet2_le2__le1346 A0 A1 Le A' A'' A0 C Le A0 A'
            Bet A0 A' A'' assms(7) assms(8) cong_reflexivity l5_6)
      moreover
      have " A. Bet A0 A A''  A PO A'' CongA A0 PO A1  A0 A1 Le A A''" 
      proof -
        have "Bet A0 A' A''" 
          using Bet A0 A' A'' by fastforce
        moreover have "A' PO A'' CongA A0 PO A1" 
          by (meson not_conga not_conga_sym A PO A' CongA A' PO A'' 
              A PO A' CongA A0 PO A1)
        moreover have "A0 A1 Le A' A''" 
          using A0 A1 Le A' A'' by auto
        ultimately show ?thesis by blast
      qed
      ultimately
      have " P Q R. (GradA A0 PO A1 P Q R  (A0 PO B LeA P Q R  
                  ( A'. Bet A0 A1 A'  Bet A0 A' B  P Q R CongA A0 PO A'  A0 C' Le A0 A' 
                     ( A. Bet A0 A A'  A PO A' CongA A0 PO A1  A0 A1 Le A A'))))" 
        using GradA A0 PO A1 P' Q' R' by blast
    }
    ultimately
    have " P Q R. (GradA A0 PO A1 P Q R  (A0 PO B LeA P Q R  
                  ( A'. Bet A0 A1 A'  Bet A0 A' B  P Q R CongA A0 PO A'  A0 C' Le A0 A' 
                     ( A. Bet A0 A A'  A PO A' CongA A0 PO A1  A0 A1 Le A A'))))" 
      using GradA A0 PO A1 P Q R by blast
  }
  ultimately
  show ?thesis 
    using P2 by blast
qed

lemma acute_archi_aux2:
  assumes "Per PO A0 B" and
    "PO  A0" and
    "B  A0" and
    "Bet A0 A1 B" and
    "A0  A1" and
    "Grad A0 A1 C"
  shows " P Q R. (GradA A0 PO A1 P Q R  (A0 PO B LeA P Q R 
                  ( A'. Bet A0 A1 A'  Bet A0 A' B  P Q R CongA A0 PO A'  A0 C Le A0 A' 
                  ( A. Bet A0 A A'  A PO A' CongA A0 PO A1  A0 A1 Le A A'))))" 
proof -
  have "¬ Col PO A0 B" 
    by (metis assms(1) assms(2) assms(3) col_permutation_1 l8_8 per_col)
  have "¬ Col A0 PO A1" 
    by (metis ¬ Col PO A0 B assms(4) assms(5) bet_col col_trivial_3 colx not_col_permutation_2)
  have "PO  A1" 
    using ¬ Col A0 PO A1 assms(4) bet_col1 by blast
  have "PO  B" 
    using assms(1) assms(3) per_distinct_1 by auto
  have "GradI A0 A1 C" 
    by (simp add: Grad__GradI assms(6))
  let ?th = " P Q R. (GradA A0 PO A1 P Q R  (A0 PO B LeA P Q R 
                  ( A'. Bet A0 A1 A'  Bet A0 A' B  P Q R CongA A0 PO A'  A0 C Le A0 A' 
                  ( A. Bet A0 A A'  A PO A' CongA A0 PO A1  A0 A1 Le A A'))))" 
  have ?th
  proof (rule GradI.induct [where ?A="A0" and ?B="A1" and ?x="C"])
    show "GradI A0 A1 C" 
      using GradI A0 A1 C by blast
    show "P Q R.
       GradA A0 PO A1 P Q R  (A0 PO B LeA P Q R 
        (A'. Bet A0 A1 A'  Bet A0 A' B  P Q R CongA A0 PO A' 
              A0 A1 Le A0 A'  (A. Bet A0 A A'  A PO A' CongA A0 PO A1  A0 A1 Le A A')))" 
      by (metis bet_col assms(1) assms(2) assms(3) assms(4) assms(5) cong2_per2__cong_conga2 
          cong_reflexivity grada_init le_reflexivity not_bet_distincts
          not_col_permutation_5 per_col)
    {
      fix C0 C'
      assume H1: "GradI A0 A1 C0" and
        H2: "P Q R. GradA A0 PO A1 P Q R  (A0 PO B LeA P Q R 
         (A'. Bet A0 A1 A'  Bet A0 A' B  P Q R CongA A0 PO A' 
           A0 C0 Le A0 A'  (A. Bet A0 A A'  A PO A' CongA A0 PO A1  A0 A1 Le A A')))" and
        H3: "Bet A0 C0 C'" and
        H4: "Cong A0 A1 C0 C'" 
      have "P Q R. GradA A0 PO A1 P Q R  (A0 PO B LeA P Q R 
           (A'. Bet A0 A1 A'  Bet A0 A' B  P Q R CongA A0 PO A' 
                 A0 C' Le A0 A'  (A. Bet A0 A A'  A PO A' CongA A0 PO A1  A0 A1 Le A A')))"
      proof (rule acute_archi_aux2_2 [where ?C="C0"])
        show "Per PO A0 B"
          using assms(1) by auto
        show "PO  A0" 
          by (simp add: assms(2))
        show "B  A0" 
          by (simp add: assms(3))
        show "Bet A0 A1 B" 
          by (simp add: assms(4))
        show "A0  A1"
          by (simp add: assms(5))
        show "Grad A0 A1 C0" 
          by (simp add: Grad_GradI H1)
        show "Bet A0 C0 C'" 
          using H3 by auto
        show "Cong A0 A1 C0 C'" 
          using H4 by auto
        show "¬ Col PO A0 B" 
          using ¬ Col PO A0 B by auto
        show "¬ Col A0 PO A1" 
          by (simp add: ¬ Col A0 PO A1)
        show "PO  A1" 
          by (simp add: PO  A1)
        show "PO  B" 
          using PO  B by auto
        show "Per PO A0 B  PO  A0  B  A0  Bet A0 A1 B  A0  A1  
               ¬ Col PO A0 B  ¬ Col A0 PO A1  PO  A1 
               (P Q R. GradA A0 PO A1 P Q R  (A0 PO B LeA P Q R 
                  (A'. Bet A0 A1 A'  Bet A0 A' B  P Q R CongA A0 PO A' 
               A0 C0 Le A0 A'  (A. Bet A0 A A'  A PO A' CongA A0 PO A1  A0 A1 Le A A'))))" 
          using H2 by blast
      qed
    }
    thus "C C'.
       GradI A0 A1 C 
       P Q R. GradA A0 PO A1 P Q R  (A0 PO B LeA P Q R 
           (A'. Bet A0 A1 A'  Bet A0 A' B  P Q R CongA A0 PO A' 
                 A0 C Le A0 A'  (A. Bet A0 A A'  A PO A' CongA A0 PO A1  A0 A1 Le A A'))) 
       Bet A0 C C' 
       Cong A0 A1 C C' 
       P Q R.
          GradA A0 PO A1 P Q R 
          (A0 PO B LeA P Q R 
           (A'. Bet A0 A1 A' 
                 Bet A0 A' B 
                 P Q R CongA A0 PO A' 
                 A0 C' Le A0 A'  (A. Bet A0 A A'  A PO A' CongA A0 PO A1  A0 A1 Le A A')))" 
      by blast
  qed
  thus ?thesis by blast
qed

lemma archi_in_acute_angles:
  assumes "archimedes_axiom" 
  shows " A B C D E F. ¬ Col A B C  Acute D E F 
                         ( P Q R. GradA A B C P Q R  D E F LeA P Q R)" 
proof -
  {
    fix A B C D E F
    assume "¬ Col A B C" and
      "Acute D E F"
    have "A  B"  
      using ¬ Col A B C col_trivial_1 by fastforce
    have "C  B" 
      using ¬ Col A B C col_trivial_2 by force
    have "E  D" 
      using Acute D E F acute_distincts by blast
    have "E  F" 
      using Acute D E F acute_distincts by blast
    have " P Q R. GradA A B C P Q R  D E F LeA P Q R" 
    proof (cases "Col D E F")
      case True
      thus ?thesis 
        by (metis Col_def Out_cases bet_out bet_out_1 l11_31_1  A  B 
            Acute D E F C  B E  D E  F acute_not_bet grada_ABC)
    next
      case False
      {
        assume "D E F LeA A B C"
        hence " P Q R. GradA A B C P Q R  D E F LeA P Q R" 
          by (metis A  B C  B grada_ABC)
      }
      moreover
      {
        assume "A B C LeA D E F"
        obtain D0 where "Col D E D0" and "D E Perp F D0" 
          using False l8_18_existence by blast
        have "E Out D0 D" 
        proof (rule acute_col_perp__out [where ?A="F"])
          show "Acute F E D"             
            using Acute D E F acute_sym by blast
          show "Col E D D0" 
            using Col D E D0 not_col_permutation_4 by blast
          show "E D Perp F D0" 
            using D E Perp F D0 perp_left_comm by blast
        qed
        have "D0  F" 
          using D E Perp F D0 perp_not_eq_2 by blast
        have "D0  E" 
          using E Out D0 D out_diff1 by auto
        have "D E F CongA D0 E F" 
          by (metis acute_col_perp__out acute_sym out2__conga out_trivial
              perp_left_comm Acute D E F Col D E D0 D E Perp F D0 E  F 
              not_col_permutation_4)
        have "Acute D0 E F" 
          by (meson acute_conga__acute Acute D E F D E F CongA D0 E F)
        have "A B C LeA D0 E F" 
          by (meson conga__lea lea_trans A B C LeA D E F D E F CongA D0 E F)
        have "¬ Col D E F" 
          by (simp add: False)
        have "Per E D0 F" 
          by (meson Per_cases l8_16_1 Col D E D0 D E Perp F D0 col_trivial_2)
        obtain D1' where "A B C CongA D0 E D1'" and "D0 E OS D1' F" 
          by (metis Col_cases False angle_construction_1 Col D E D0 
              D0  E ¬ Col A B C col2__eq)
        have "D0 E F CongA D0 E F" 
          using D0  E E  F conga_refl by auto
        hence "D0 E D1' LeA D0 E F" 
          using l11_30 [where ?A="A" and ?B="B" and ?C="C" and ?D="D0" and ?E="E" and ?F="F"]
            A B C LeA D0 E F A B C CongA D0 E D1' by blast
        have "D0 E OS F D1'" 
          by (simp add: D0 E OS D1' F one_side_symmetry)
        hence "D1' InAngle D0 E F" 
          using lea_in_angle D0 E D1' LeA D0 E F by blast
        then obtain D1 where "Bet D0 D1 F" and "D1 = E  E Out D1 D1'" 
          using InAngle_def by force
        have "D1 = E  ( P Q R. GradA A B C P Q R  D E F LeA P Q R)" 
          using Acute D0 E F Bet D0 D1 F acute_not_bet by blast
        moreover
        {
          assume "E Out D1 D1'"
          have "A B C CongA D0 E D1"
          proof (rule l11_10 [where ?A="A" and ?C="C" and ?D="D0" and ?F="D1'"], 
              insert A B C CongA D0 E D1' E Out D1 D1')
            show "B Out A A" 
              using A  B out_trivial by auto
            show "B Out C C" 
              by (simp add: C  B out_trivial)
            show "E Out D0 D0" 
              by (simp add: D0  E out_trivial)
          qed
          have "¬ Col D0 E D1'" 
            using D0 E OS D1' F col123__nos by force
          have "D0  D1" 
            using Col_cases E Out D1 D1' ¬ Col D0 E D1' out_col by blast
          obtain F' where "Bet D0 F F'" and "Cong F F' D0 F" 
            using segment_construction by blast
          obtain G where "Grad D0 D1 G" and "D0 F' Le D0 G" 
            using Reach_def D0  D1 archimedes_axiom_def assms by blast
          have "GradI D0 D1 G" 
            by (simp add: Grad__GradI Grad D0 D1 G)
          have " P Q R. (GradA D0 E D1 P Q R  (D0 E F LeA P Q R 
            ( A'. Bet D0 D1 A'  Bet D0 A' F  P Q R CongA D0 E A' 
              D0 G Le D0 A'  ( A. Bet D0 A A'  A E A' CongA D0 E D1  D0 D1 Le A A'))))" 
            using acute_archi_aux2 Bet D0 D1 F D0  D1 D0  E D0  F 
              Grad D0 D1 G Per E D0 F by blast
          then obtain P Q R where
            "GradA D0 E D1 P Q R" and
            "D0 E F LeA P Q R  ( A'. Bet D0 D1 A'  Bet D0 A' F 
                 P Q R CongA D0 E A'  D0 G Le D0 A'  
                 ( A. Bet D0 A A'  A E A' CongA D0 E D1  D0 D1 Le A A'))" by blast
          have "D0  E"  
            using grada_distincts [where ?A="D0" and ?B="E" and ?C="D1" and 
                ?D="P" and ?E="Q" and ?F="R"] GradA D0 E D1 P Q R by blast
          have "D1  E" 
            using grada_distincts [where ?A="D0" and ?B="E" and ?C="D1" and 
                ?D="P" and ?E="Q" and ?F="R"] GradA D0 E D1 P Q R by blast
          have "P  Q"
            using grada_distincts [where ?A="D0" and ?B="E" and ?C="D1" and 
                ?D="P" and ?E="Q" and ?F="R"] GradA D0 E D1 P Q R by blast
          have "R  Q" 
            using grada_distincts [where ?A="D0" and ?B="E" and ?C="D1" and 
                ?D="P" and ?E="Q" and ?F="R"] GradA D0 E D1 P Q R by blast
          have "GradA A B C P Q R"
          proof (rule conga2_grada__grada [where ?A="D0" and ?B="E" and ?C="D1" 
                and ?D="P" and ?E="Q" and ?F="R"], insert GradA D0 E D1 P Q R)
            show "D0 E D1 CongA A B C" 
              using A B C CongA D0 E D1 conga_sym_equiv by auto
            show "P Q R CongA P Q R" 
              by (simp add: P  Q R  Q conga_refl)
          qed
          moreover 
          have "D E F LeA P Q R" 
          proof -
            {
              assume "D0 E F LeA P Q R"
              have "D E F LeA P Q R" 
              proof (rule l11_30 [where ?A="D0" and ?B="E" and ?C="F" and 
                    ?D="P" and ?E="Q" and ?F="R"])
                show "D0 E F LeA P Q R" 
                  by (simp add: D0 E F LeA P Q R)
                show "D0 E F CongA D E F" 
                  using D E F CongA D0 E F conga_sym_equiv by auto
                show "P Q R CongA P Q R" 
                  using P  Q R  Q conga_refl by auto
              qed
            }
            moreover
            {
              assume " A'. Bet D0 D1 A'  Bet D0 A' F  P Q R CongA D0 E A' 
                D0 G Le D0 A'  ( A. Bet D0 A A'  A E A' CongA D0 E D1  D0 D1 Le A A')"
              then obtain A' where "Bet D0 D1 A'" and "Bet D0 A' F" and
                "P Q R CongA D0 E A'" and "D0 G Le D0 A'" and 
                " A. Bet D0 A A'  A E A' CongA D0 E D1  D0 D1 Le A A'"
                by blast
              have "D0 A' Le D0 F"         
                by (simp add: Bet D0 A' F bet__le1213)
              hence "D0 G Le D0 F"
                using D0 G Le D0 A' le_transitivity by blast
              moreover have "D0 F Lt D0 F'" 
                using Bet D0 F F' Cong F F' D0 F D0  F 
                  bet__lt1213 cong_diff_3 by presburger
              ultimately have "D0 G Lt D0 F'" 
                using le1234_lt__lt by blast
              hence False 
                using D0 F' Le D0 G le__nlt by auto
              hence "D E F LeA P Q R" 
                by blast
            }
            ultimately show ?thesis 
              using D0 E F LeA P Q R  (A'. Bet D0 D1 A'  Bet D0 A' F  
                P Q R CongA D0 E A'  D0 G Le D0 A'  
                (A. Bet D0 A A'  A E A' CongA D0 E D1  D0 D1 Le A A')) by blast
          qed
          ultimately have " P Q R. GradA A B C P Q R  D E F LeA P Q R" 
            by blast
        }
        ultimately have " P Q R. GradA A B C P Q R  D E F LeA P Q R" 
          using D1 = E  E Out D1 D1' by blast
      }
      ultimately show ?thesis using lea_total 
        by (metis A  B C  B E  D E  F)
    qed
  }
  thus ?thesis 
    by blast
qed

lemma angles_archi_aux:
  assumes "GradA A B C D E F" and 
    "GradA A B C G H I" and
    "¬ SAMS D E F G H I" 
  shows " P Q R. GradA A B C P Q R  ¬ SAMS P Q R A B C" 
proof -
  have "¬ SAMS D E F G H I  ( P Q R. GradA A B C P Q R  ¬ SAMS P Q R A B C)" 
  proof (rule GradA.induct [OF assms(2)])
    show "Da Ea Fa.
       A B C CongA Da Ea Fa  ¬ SAMS D E F Da Ea Fa  
                                (P Q R. GradA A B C P Q R  ¬ SAMS P Q R A B C)" 
      by (metis Tarski_neutral_dimensionless.conga_refl Tarski_neutral_dimensionless_axioms 
          assms(1) conga2_sams__sams grada_distincts)
    {
      fix Da Ea Fa G H I
      assume "GradA A B C Da Ea Fa" and
        "¬ SAMS D E F Da Ea Fa  ( P Q R. GradA A B C P Q R  ¬ SAMS P Q R A B C)" and
        "SAMS Da Ea Fa A B C" and
        "Da Ea Fa A B C SumA G H I"
      {
        assume "¬ SAMS D E F G H I"
        {
          assume "SAMS D E F Da Ea Fa"
          have "E  D" 
            using SAMS D E F Da Ea Fa sams_distincts by blast
          have "E  F" 
            using SAMS D E F Da Ea Fa sams_distincts by blast
          have "Ea  Da" 
            using SAMS D E F Da Ea Fa sams_distincts by blast
          have "Ea  Fa" 
            using SAMS D E F Da Ea Fa sams_distincts by blast
          obtain P Q R where "D E F Da Ea Fa SumA P Q R"
            using ex_suma E  D E  F Ea  Da Ea  Fa by presburger
          have "GradA A B C P Q R" 
            using grada2_sams_suma__grada [where ?D="D" and ?E="E" and ?F="F" and
                ?G="Da" and ?H="Ea" and ?I="Fa"]
              assms(1) GradA A B C Da Ea Fa SAMS D E F Da Ea Fa D E F Da Ea Fa SumA P Q R
            by blast
          moreover 
          {
            assume "SAMS P Q R A B C"
            have "SAMS D E F G H I"
              using SAMS D E F Da Ea Fa SAMS Da Ea Fa A B C D E F Da Ea Fa SumA P Q R
                Da Ea Fa A B C SumA G H I SAMS P Q R A B C sams_assoc_1 by blast
            hence False 
              using ¬ SAMS D E F G H I by blast
          }
          ultimately have "P Q R. GradA A B C P Q R  ¬ SAMS P Q R A B C" 
            by blast
        }
        hence "P Q R. GradA A B C P Q R  ¬ SAMS P Q R A B C" 
          using ¬ SAMS D E F Da Ea Fa  (P Q R. GradA A B C P Q R  ¬ SAMS P Q R A B C) 
          by blast
      }
      hence "¬ SAMS D E F G H I  (P Q R. GradA A B C P Q R  ¬ SAMS P Q R A B C)" 
        by blast
    }
    thus "Da Ea Fa G H I.
       GradA A B C Da Ea Fa 
       ¬ SAMS D E F Da Ea Fa  (P Q R. GradA A B C P Q R  ¬ SAMS P Q R A B C) 
       SAMS Da Ea Fa A B C  Da Ea Fa A B C SumA G H I  
       ¬ SAMS D E F G H I  (P Q R. GradA A B C P Q R  ¬ SAMS P Q R A B C)" 
      by blast
  qed
  thus ?thesis 
    using assms(3) by blast
qed

lemma angles_archi_aux1:
  assumes "archimedes_axiom"
  shows " A B C D E F.
¬ Col A B C  ¬ Bet D E F 
( P Q R. GradA A B C P Q R  (D E F LeA P Q R  ¬ SAMS P Q R A B C))" 
proof -
  {
    fix A B C D E F
    assume "¬ Col A B C" and
      "¬ Bet D E F"
    have "D  E" 
      using ¬ Bet D E F between_trivial2 by auto
    have "F  E" 
      using ¬ Bet D E F not_bet_distincts by blast
    obtain F1 where "F1 InAngle D E F" and "F1 E D CongA F1 E F" 
      using angle_bisector D  E F  E by blast
    have "F1  E" 
      using F1 E D CongA F1 E F conga_distinct by auto
    have "¬ E F1 OS D F"
    proof (cases "Col D E F1")
      case True
      thus ?thesis
        using NCol_perm col123__nos by blast
    next
      case False
      thus ?thesis
        by (meson F1 InAngle D E F col124__nos col_permutation_4 col_permutation_5 
            in_angle_two_sides invert_one_side l9_9)
    qed
    have "D E F1 D E F1 SumA D E F" 
      by (meson conga_refl D  E F  E F1 E D CongA F1 E F 
          F1 InAngle D E F F1  E conga3_suma__suma 
          conga_left_comm inangle__suma not_conga_sym)
    have "SAMS D E F1 D E F1" 
    proof -
      {
        assume "Bet D E F1" 
        hence False 
          using bet_in_angle_bet F1 InAngle D E F ¬ Bet D E F by blast
      }
      hence "E Out D F1  ¬ Bet D E F1" 
        by blast
      moreover 
      have " J. F1 E J CongA D E F1  ¬ E F1 OS D J  ¬ D E TS F1 J  Coplanar D E F1 J" 
      proof -
        have "F1 E F CongA D E F1" 
          using F1 E D CongA F1 E F conga_left_comm conga_sym_equiv by blast
        moreover have "¬ E F1 OS D F" 
          by (simp add: ¬ E F1 OS D F)
        moreover have "¬ D E TS F1 F" 
        proof (cases "Col D E F1")
          case True
          thus ?thesis 
            using TS_def col_permutation_2 by blast
        next
          case False
          thus ?thesis
            by (metis Col_cases TS_def F1 InAngle D E F in_angle_one_side l9_9)
        qed
        moreover have "Coplanar D E F1 F" 
          by (meson inangle__coplanar F1 InAngle D E F coplanar_perm_8)
        ultimately show ?thesis 
          by blast
      qed
      ultimately show ?thesis 
        using SAMS_def D  E by auto
    qed
    hence "Acute D E F1" 
      by (metis nbet_sams_suma__acute D E F1 D E F1 SumA D E F ¬ Bet D E F)
    then obtain P1 Q1 R1 where "GradA A B C P1 Q1 R1" and "D E F1 LeA P1 Q1 R1"  
      using archi_in_acute_angles ¬ Col A B C assms by blast
    have "P1  Q1"
      using GradA A B C P1 Q1 R1 grada_distincts by blast
    have "Q1  R1" 
      using GradA A B C P1 Q1 R1 grada_distincts by blast
    {
      assume "SAMS P1 Q1 R1 P1 Q1 R1"
      obtain P Q R where "P1 Q1 R1 P1 Q1 R1 SumA P Q R" 
        using ex_suma P1  Q1 Q1  R1 by presburger
      have "GradA A B C P Q R" 
        using grada2_sams_suma__grada [where ?D="P1" and ?E="Q1" and ?F="R1" and 
            ?G="P1" and ?H="Q1" and ?I="R1"] GradA A B C P1 Q1 R1 GradA A B C P1 Q1 R1 
          SAMS P1 Q1 R1 P1 Q1 R1 P1 Q1 R1 P1 Q1 R1 SumA P Q R by blast
      moreover have "D E F LeA P Q R" 
        using sams_lea2_suma2__lea [where ?A="D" and ?B="E" and ?C="F1" and 
            ?D="D" and ?E="E" and ?F="F1" and
            ?A'="P1" and ?B'="Q1" and ?C'="R1" and ?D'="P1" and ?E'="Q1" and ?F'="R1"]
          D E F1 LeA P1 Q1 R1 SAMS P1 Q1 R1 P1 Q1 R1 D E F1 D E F1 SumA D E F
          P1 Q1 R1 P1 Q1 R1 SumA P Q R by blast
      ultimately have " P Q R. GradA A B C P Q R  (D E F LeA P Q R  ¬ SAMS P Q R A B C)" 
        by blast
    }
    moreover
    {
      assume "¬ SAMS P1 Q1 R1 P1 Q1 R1"
      hence " P Q R. GradA A B C P Q R  (D E F LeA P Q R  ¬ SAMS P Q R A B C)" 
        using angles_archi_aux GradA A B C P1 Q1 R1 calculation by blast
    }
    ultimately have " P Q R. GradA A B C P Q R  (D E F LeA P Q R  ¬ SAMS P Q R A B C)" 
      by blast
  }
  thus ?thesis
    by blast
qed

(** Inspired by Hartshorne's demonstration of Lemma 35.1 in Geometry Euclid and Beyond *)
lemma archi_in_angles:
  assumes "archimedes_axiom" 
  shows " A B C.  D ::'p.  E ::'p.  F ::'p. (¬ Col A B C  D  E  F  E) 
    ( P Q R. GradA A B C P Q R  (D E F LeA P Q R  ¬ SAMS P Q R A B C))" 
proof -
  {
    fix A B C 
    fix D::'p
    fix E::'p
    fix F::'p
    assume "¬ Col A B C" and 
      "D  E" and 
      "F  E"
    have " P Q R. (GradA A B C P Q R  (D E F LeA P Q R  ¬ SAMS P Q R A B C))" 
    proof (cases "Bet D E F")
      case True
      obtain A0 where "Bet A B A0" and "Cong B A0 A B" 
        using segment_construction by blast
      have "A  B" 
        using ¬ Col A B C not_col_distincts by blast
      have "C  B" 
        using ¬ Col A B C col_trivial_2 by auto
      have "B  A0" 
        using A  B Cong B A0 A B cong_reverse_identity by blast
      have "¬ Col A0 B C" 
        by (meson B  A0 Bet A B A0 ¬ Col A B C bet_col col2__eq col_permutation_3)
      obtain P1 Q1 R1 where "GradA A B C P1 Q1 R1" and 
        "C B A0 LeA P1 Q1 R1  ¬ SAMS P1 Q1 R1 A B C" 
        using angles_archi_aux1 
        by (metis Col_def ¬ Col A B C ¬ Col A0 B C assms between_symmetry)
      {
        assume "SAMS P1 Q1 R1 A B C"
        hence "C B A0 LeA P1 Q1 R1" 
          using C B A0 LeA P1 Q1 R1  ¬ SAMS P1 Q1 R1 A B C by auto
        have "P1  Q1" 
          using SAMS P1 Q1 R1 A B C 
          by (simp add: sams_distincts)
        have "R1  Q1" 
          using C B A0 LeA P1 Q1 R1 lea_distincts by blast
        obtain P Q R where "P1 Q1 R1 A B C SumA P Q R" 
          using ex_suma A  B C  B P1  Q1 R1  Q1 by presburger
        have "GradA A B C P Q R"
          using grada_stab [where ?D="P1" and ?E="Q1" and ?F="R1"]
            GradA A B C P1 Q1 R1 SAMS P1 Q1 R1 A B C P1 Q1 R1 A B C SumA P Q R by auto
        moreover 
        have "P  Q" 
          using calculation grada_distincts by blast
        have "R  Q" 
          using calculation grada_distincts by blast
        have "A B A0 LeA P Q R"
        proof (rule sams_lea2_suma2__lea [where 
              ?A="A0" and ?B="B" and ?C="C" and ?D="A" and ?E="B" and ?F="C" 
              and ?A'="P1" and ?B'="Q1" and ?C'="R1" and ?D'="A" and ?E'="B" and ?F'="C"])
          show "A0 B C LeA P1 Q1 R1" 
            using C B A0 LeA P1 Q1 R1 lea_left_comm by blast
          show "A B C LeA A B C" 
            using A  B C  B lea_refl by force
          show "SAMS P1 Q1 R1 A B C" 
            using SAMS P1 Q1 R1 A B C by auto
          show "A0 B C A B C SumA A B A0" 
            by (metis Bet_cases A  B B  A0 Bet A B A0 C  B 
                bet__suma suma_middle_comm suma_right_comm)
          show "P1 Q1 R1 A B C SumA P Q R" 
            using P1 Q1 R1 A B C SumA P Q R by auto
        qed
        hence "Bet P Q R" 
          using Bet A B A0 bet_lea__bet by blast
        hence "D E F LeA P Q R" 
          using l11_31_2 D  E F  E P  Q R  Q by blast
        ultimately have " P Q R. (GradA A B C P Q R  (D E F LeA P Q R  ¬ SAMS P Q R A B C))" 
          by blast
      }
      moreover
      {
        assume "¬ SAMS P1 Q1 R1 A B C"
        hence " P Q R. (GradA A B C P Q R  (D E F LeA P Q R  ¬ SAMS P Q R A B C))" 
          using GradA A B C P1 Q1 R1 by blast
      }
      ultimately show ?thesis 
        by blast
    next
      case False
      thus ?thesis 
        using ¬ Col A B C angles_archi_aux1 assms by blast
    qed
  }
  thus ?thesis
    by blast
qed

(** If Archimedes' postulate holds, every nondegenerate angle can be 
    repeated until exceeding 180° *)
lemma archi__grada_destruction:
  assumes "archimedes_axiom" 
  shows " A B C. ¬ Col A B C 
( P Q R. GradA A B C P Q R  ¬ SAMS P Q R A B C)" 
proof -
  {
    fix A B C
    assume "¬ Col A B C"
    obtain A0 where "Bet A B A0" and "Cong B A0 A B" 
      using segment_construction by blast
    have "A  B" 
      using ¬ Col A B C col_trivial_1 by blast
    have "B  A0" 
      using A  B Cong B A0 A B cong_reverse_identity by blast
    obtain P Q R where "GradA A B C P Q R" and "A B A0 LeA P Q R  ¬ SAMS P Q R A B C"
      using archi_in_angles A  B B  A0 ¬ Col A B C assms by metis
    {
      assume "A B A0 LeA P Q R"
      assume "SAMS P Q R A B C"
      hence "B Out A C  ¬ Bet P Q R"
        using SAMS_def by blast
      have "B Out A C  False" 
        using Col_cases ¬ Col A B C out_col by blast
      moreover
      {
        assume "¬ Bet P Q R"
        have "Bet A B A0" 
          by (simp add: Bet A B A0)
        hence False 
          using bet_lea__bet A B A0 LeA P Q R B Out A C  ¬ Bet P Q R calculation by blast
      }
      hence "¬ Bet P Q R  False" 
        by blast
      ultimately have False 
        using B Out A C  ¬ Bet P Q R by fastforce
    }
    hence "A B A0 LeA P Q R  ¬ SAMS P Q R A B C" 
      by blast
    hence "¬ SAMS P Q R A B C" 
      using A B A0 LeA P Q R  ¬ SAMS P Q R A B C by blast
    hence " P Q R. GradA A B C P Q R  ¬ SAMS P Q R A B C" 
      using GradA A B C P Q R by blast
  }
  thus ?thesis 
    by blast
qed

lemma gradaexp_destruction_aux:
  assumes "GradA A B C P Q R"
  shows " S T U. GradAExp A B C S T U  (Obtuse S T U  P Q R LeA S T U)" 
proof (rule GradA.induct [OF assms(1)])
  show "D E F. A B C CongA D E F  
           S T U. GradAExp A B C S T U  (Obtuse S T U  D E F LeA S T U)" 
    by (metis conga__lea456123 conga_diff1 conga_diff2 gradaexp_ABC)
  {
    fix D E F G H I
    assume "GradA A B C D E F" and
      "S T U. GradAExp A B C S T U  (Obtuse S T U  D E F LeA S T U)" and
      "SAMS D E F A B C" and
      "D E F A B C SumA G H I"
    then obtain P Q R where "GradAExp A B C P Q R" and "Obtuse P Q R  D E F LeA P Q R" 
      by blast
    have "P  Q" 
      using GradAExp A B C P Q R gradaexp_distincts by blast
    have "R  Q" 
      using GradAExp A B C P Q R gradaexp_distincts by blast
    {
      assume "SAMS P Q R P Q R"
      {
        assume "Obtuse P Q R" 
        hence "S T U. GradAExp A B C S T U  (Obtuse S T U  G H I LeA S T U)" 
          using GradAExp A B C P Q R by blast
      }
      moreover
      {
        assume  "D E F LeA P Q R" 
        obtain S T U where "P Q R P Q R SumA S T U" 
          using ex_suma P  Q R  Q by presburger
        have "GradAExp A B C S T U"
        proof (rule gradaexp_stab [where ?D="P" and ?E="Q" and ?F="R"])
          show "GradAExp A B C P Q R" 
            by (simp add: GradAExp A B C P Q R)
          show "SAMS P Q R P Q R" 
            by (simp add: SAMS P Q R P Q R)
          show "P Q R P Q R SumA S T U"
            using P Q R P Q R SumA S T U by auto
        qed
        moreover
        have "GradA A B C P Q R" 
          using GradAExp A B C P Q R gradaexp__grada by auto
        hence "G H I LeA S T U" 
          by (meson D E F A B C SumA G H I D E F LeA P Q R P Q R P Q R SumA S T U
              SAMS P Q R P Q R grada__lea sams_lea2_suma2__lea)
        ultimately have "S T U. GradAExp A B C S T U  (Obtuse S T U  G H I LeA S T U)" 
          by blast
      }
      ultimately have "S T U. GradAExp A B C S T U  (Obtuse S T U  G H I LeA S T U)" 
        using Obtuse P Q R  D E F LeA P Q R by blast
    }
    moreover
    {
      assume "¬ SAMS P Q R P Q R"
      hence "Obtuse P Q R" 
        using P  Q R  Q nsams__obtuse by auto
      hence "S T U. GradAExp A B C S T U  (Obtuse S T U  G H I LeA S T U)" 
        using GradAExp A B C P Q R by blast
    }
    ultimately have "S T U. GradAExp A B C S T U  (Obtuse S T U  G H I LeA S T U)" 
      by blast
  }
  thus "D E F G H I.
       GradA A B C D E F 
       S T U. GradAExp A B C S T U  (Obtuse S T U  D E F LeA S T U) 
       SAMS D E F A B C 
       D E F A B C SumA G H I  S T U. GradAExp A B C S T U  (Obtuse S T U  G H I LeA S T U)" 
    by blast
qed

(** If Archimedes' postulate holds, every nondegenerate angle can be doubled until exceeding 90° *)
lemma archi__gradaexp_destruction:
  assumes "archimedes_axiom"
  shows " A B C. ¬ Col A B C  ( P Q R. GradAExp A B C P Q R  Obtuse P Q R)" 
proof -
  {
    fix A B C
    assume "¬ Col A B C"
    obtain D E F where "GradA A B C D E F" and "¬ SAMS D E F A B C" 
      using archi__grada_destruction ¬ Col A B C assms by blast
    obtain P Q R where "GradAExp A B C P Q R" and "Obtuse P Q R  D E F LeA P Q R" 
      using GradA A B C D E F gradaexp_destruction_aux by blast
    have "P  Q" 
      using GradAExp A B C P Q R gradaexp_distincts by blast
    have "R  Q" 
      using GradAExp A B C P Q R gradaexp_distincts by blast
    {
      assume "D E F LeA P Q R" 
      {
        assume "SAMS P Q R P Q R" 
        have "A B C LeA P Q R" 
          using GradAExp A B C P Q R grada__lea gradaexp__grada by blast
        hence "SAMS D E F A B C" 
          using sams_lea2__sams [where ?A'="P" and ?B'="Q" and ?C'="R" and 
              ?D'="P" and ?E'="Q" and ?F'="R"]
            SAMS P Q R P Q R D E F LeA P Q R by blast
        hence False 
          using ¬ SAMS D E F A B C by blast
      }
      hence "¬ SAMS P Q R P Q R" 
        by blast
      hence "Obtuse P Q R" 
        using P  Q R  Q nsams__obtuse by auto
    }
    hence " P Q R. GradAExp A B C P Q R  Obtuse P Q R" 
      using GradAExp A B C P Q R Obtuse P Q R  D E F LeA P Q R by blast
  }
  thus ?thesis
    by blast
qed

subsection "Equivalence Greenberg - Aristotle"

lemma aristotle__greenberg: 
  assumes "aristotle_s_axiom" 
  shows "greenberg_s_axiom" 
proof -
  {
    fix P Q R A B C
    assume "¬ Col A B C" and
      "Acute A B C" and "Q  R" and "Per P Q R"
    have " S. P S Q LtA A B C  Q Out S R" 
    proof (cases "P = Q")
      case True
      have "P R P LtA A B C" 
      proof -
        have "P R P LeA A B C" 
          using True Q  R ¬ Col A B C 
            lea121345 not_col_distincts by presburger
        moreover
        have "¬ P R P CongA A B C" 
          using ¬ Col A B C col_conga_col col_trivial_3 by blast
        ultimately
        show ?thesis 
          by (simp add: LtA_def)
      qed
      moreover
      have "P Out R R" 
        using True Q  R out_trivial by auto
      ultimately
      show ?thesis 
        using True by blast
    next
      case False
      have " X Y. B Out A X  B Out C Y  Per B X Y  P Q Lt X Y" 
        using aristotle_s_axiom_def ¬ Col A B C 
          Acute A B C assms by blast
      then obtain X Y where "B Out A X" and "B Out C Y" and 
        "Per B X Y" and "P Q Lt X Y" 
        by blast
      have "P Q Le X Y" 
        by (simp add: P Q Lt X Y lt__le)
      have "¬ Cong P Q X Y" 
        using P Q Lt X Y cong__nlt by blast
      have "X  Y" 
        using P Q Lt X Y lt_diff by presburger
      have "P  Q" 
        using False by blast
      hence " T. Q Out T P  Cong Q T X Y" 
        using l6_11_existence X  Y by simp
      then obtain T where "Q Out T P" and "Cong Q T X Y"
        by blast 
      have "X  B" 
        using B Out A X out_diff2 by blast
      have "R  Q" 
        using Q  R by auto
      hence " S. Q Out S R  Cong Q S X B" 
        using l6_11_existence R  Q X  B by simp
      then obtain S where "Q Out S R" and "Cong Q S X B"
        by blast 
      have "Per S Q P" 
      proof -
        have "Per R Q P" 
          using Per P Q R l8_2 by blast
        moreover
        have "Col Q R S" 
          using Q Out S R l6_6 out_col by blast
        ultimately
        show ?thesis 
          using l8_3 R  Q by blast
      qed
      have "Per T Q S" 
      proof -
        have "Per P Q S" 
          by (simp add: Per S Q P l8_2)
        moreover
        have" Col Q P T" 
          by (simp add: Q Out T P l6_6 out_col)
        ultimately
        show ?thesis 
          by (metis False l8_2 per_col)
      qed
      have "P  S" 
        using Per S Q P Q Out S R l6_3_1 l8_8 by blast
      have "T  S" 
        using Per T Q S Q Out T P l6_3_1 l8_8 by blast
      have "P S Q LtA A B C" 
      proof -
        have "P S Q CongA P S Q" 
          by (metis P  S Q Out S R conga_refl out_diff1)
        moreover
        have "T S Q CongA A B C" 
        proof -
          have "T S Q CongA X B Y" 
            by (metis Per_cases cong_diff_3 Cong Q S X B 
                Cong Q T X Y Per B X Y Per T Q S Q Out T P 
                T  S X  B conga_right_comm l10_12 
                l11_51 l6_3_1 not_cong_2143)
          moreover
          have "X B Y CongA A B C" 
            by (simp add: B Out A X B Out C Y out2__conga)
          ultimately
          show ?thesis 
            using not_conga by blast
        qed
        moreover
        have "P S Q LtA T S Q" 
        proof -
          have "Q S P LeA Q S T"
          proof -
            have "P InAngle Q S T" 
            proof -
              have "Bet Q P T" 
              proof -
                have "Q Out P T" 
                  by (simp add: Q Out T P l6_6)
                moreover
                have "Q P Le Q T" 
                proof -
                  have "Q P Le X Y" 
                    using Le_cases P Q Le X Y by blast
                  moreover
                  have "X Y Le Q T" 
                    by (simp add: Cong Q T X Y cong__le3412)
                  ultimately
                  show ?thesis 
                    using le_transitivity by blast
                qed
                ultimately
                show ?thesis 
                  by (simp add: l6_13_1)
              qed
              moreover
              have "P = S  S Out P P" 
                using out_trivial by blast
              ultimately
              show ?thesis 
                using InAngle_def Out_def P  S Q Out S R T  S by auto
            qed
            moreover
            have "Q S T CongA Q S T" 
              by (metis Q Out S R T  S conga_refl l6_3_1)
            ultimately
            show ?thesis 
              by (simp add: inangle__lea)
          qed
          hence "P S Q LeA T S Q" 
            by (simp add: lea_comm)
          moreover
          {
            assume "P S Q CongA T S Q" 
            have "Cong Q P Q T  Cong S P S T  Q P S CongA Q T S" 
            proof -
              {
                assume "Col Q S P"
                hence "Col S Q P" 
                  using Col_cases by blast
                hence "S = Q  P = Q" 
                  by (simp add: Per S Q P l8_9)
                hence False 
                  using False Out_def Q Out S R by blast
              }
              hence "¬ Col Q S P" 
                by blast
              moreover
              have "S Q P CongA S Q T" 
                using Q Out S R Q Out T P out2__conga 
                  out_diff1 out_trivial by force
              moreover
              have "Q S P CongA Q S T" 
                by (simp add: P S Q CongA T S Q conga_comm)
              moreover
              have "Cong Q S Q S" 
                by (simp add: cong_reflexivity)
              ultimately
              show ?thesis 
                by (metis l11_50_1)
            qed
            have "Cong P Q X Y" 
            proof -
              have "Cong P Q T Q" 
                using Cong_cases 
                  Cong Q P Q T  Cong S P S T  Q P S CongA Q T S 
                by blast
              moreover
              have "Cong T Q X Y" 
                by (meson Cong Q T X Y not_cong_2134)
              ultimately
              show ?thesis 
                using cong_transitivity by blast
            qed
            hence False 
              using ¬ Cong P Q X Y by auto
          }
          hence "¬ P S Q CongA T S Q" 
            by auto
          ultimately
          show ?thesis 
            by (simp add: LtA_def)
        qed
        ultimately
        show ?thesis 
          using conga_preserves_lta by blast
      qed
      moreover
      have "Q Out S R" 
        by (simp add: Q Out S R)
      ultimately
      show ?thesis 
        by blast
    qed
  }
  thus ?thesis 
    by (simp add: greenberg_s_axiom_def)
qed

(** This proof is inspired by The elementary Archimedean axiom in absolute geometry, 
by Victor Pambuccian *)

lemma greenberg__aristotle:
  assumes "greenberg_s_axiom" 
  shows "aristotle_s_axiom" 
proof -  
  {
    fix P Q A B C
    assume "¬ Col A B C" and "Acute A B C"
    obtain D' where "A B Perp D' B" and "A B OS C D'" 
      using ¬ Col A B C col_trivial_2 l10_15 by blast
    have " X. Col A B X  A B Perp C X" 
      using ¬ Col A B C l8_18_existence by auto
    then obtain X where "Col A B X" and "A B Perp C X" 
      by blast
    have " X Y. B Out A X  B Out C Y  Per B X Y  P Q Lt X Y" 
    proof (cases "P = Q")
      case True
      have "B Out A X" 
      proof -
        have "Acute C B A" 
          by (simp add: Acute A B C acute_sym)
        moreover
        have "Col B A X" 
          using Col A B X not_col_permutation_4 by blast
        moreover
        have "B A Perp C X" 
          by (simp add: perp_left_comm A B Perp C X)
        ultimately
        show ?thesis 
          using acute_col_perp__out l6_6 by blast
      qed
      moreover
      have "B Out C C" 
        by (metis ¬ Col A B C col_trivial_2 out_trivial)
      moreover
      have "Per B X C" 
      proof (cases "B = X")
        case True
        thus ?thesis 
          using l8_20_1_R1 by force
      next
        case False
        have "Col A B B" 
          using col_trivial_2 by blast
        thus ?thesis 
          by (metis Col A B X A B Perp C X False perp_col2 perp_per_1)
      qed
      moreover
      have "X  C" 
        using Col A B X ¬ Col A B C by auto
      hence "P P Lt X C" 
        using lt1123 by simp
      ultimately
      show ?thesis 
        using True by blast
    next
      case False
      hence "P  Q" 
        by simp
      have "B  D'" 
        using A B OS C D' os_distincts by blast
      then obtain P' where "B Out D' P'" and "Cong B P' P Q" 
        using segment_construction_3 P  Q by blast
      have " C'. P' C' B LtA A B C  B Out C' A" 
      proof -
        have "B  A" 
          using ¬ Col A B C col_trivial_1 by blast
        moreover
        have "Per P' B A" 
          by (metis Perp_cases A B Perp D' B
              B Out D' P' l8_20_1_R1 out_col
              perp_col perp_per_1)
        ultimately
        show ?thesis 
          using  ¬ Col A B C Acute A B C 
            greenberg_s_axiom_def assms by blast
      qed
      then obtain C' where "P' C' B LtA A B C" and "B Out C' A" 
        by blast
      have " D''. B C' Perp D'' C'  B C' OS C D''" 
      proof -
        have "Col B C' C'" 
          using col_trivial_2 by blast
        moreover
        have "¬ Col B C' C" 
          by (metis Out_def B Out C' A ¬ Col A B C
              col_transitivity_1 not_col_permutation_4 out_col)
        ultimately
        show ?thesis 
          by (simp add: l10_15)
      qed
      then obtain D'' where "B C' Perp D'' C'" and "B C' OS C D''" 
        by blast
      have "C'  D''" 
        using B C' Perp D'' C' perp_not_eq_2 by blast
      then obtain P'' where "C' Out D'' P''" and "Cong C' P'' P Q"
        using False segment_construction_3 by blast 
      have "B  C" 
        using ¬ Col A B C col_trivial_2 by blast
      have "B  P''" 
        by (metis Col_cases B C' OS C D'' 
            C' Out D'' P'' one_side_not_col124 out_col)
      then obtain Z where "B Out C Z" and "Cong B Z B P''" 
        using segment_construction_3 B  C by blast 
      have "¬ Col A B Z" 
        by (metis Col_cases Out_def B Out C Z
            ¬ Col A B C l6_16_1 out_col)
      then obtain Z' where "Col A B Z'" and "A B Perp Z Z'" 
        using l8_18_existence by blast
      have "B Out A Z'" 
        by (metis A B Perp Z Z' Acute A B C 
            B Out C Z B Out C' A Col A B Z' 
            acute_col_perp__out acute_out2__acute 
            acute_sym bet2__out l6_3_1 l6_6 
            not_col_permutation_4 perp_left_comm)
      moreover
      have "B Out C Z" 
        by (simp add: B Out C Z)
      moreover
      have "Per B Z' Z" 
      proof -
        have "A B Perp Z' Z" 
          using Perp_perm A B Perp Z Z' by blast
        moreover
        have "Z  B" 
          using ¬ Col A B Z col_trivial_2 by blast
        moreover
        have "Col A B Z'" 
          by (simp add: Col A B Z')
        moreover
        have "Col A B B" 
          by (simp add: col_trivial_2)
        ultimately
        show ?thesis 
          by (meson l8_16_1 A B Perp Z Z' l8_2)
      qed
      moreover
      have "C' P'' Lt Z' Z" 
      proof -
        have "Per B Z' Z" 
          by (simp add: calculation(3))
        moreover
        have "Per B C' P''" 
          by (metis out_distinct B C' Perp D'' C' 
              C' Out D'' P'' out_col perp_col1 
              perp_comm perp_per_2)
        moreover
        have "B Z' Lt B C'" 
        proof -
          have "B C' P' CongA A B P''" 
          proof -
            have "B C' P' CongA C' B P''" 
            proof -
              have "P' B C' CongA P'' C' B" 
              proof -
                have "B P' Perp C' B" 
                proof -
                  have "B D' Perp C' B" 
                  proof -
                    have "B A Perp B D'" 
                      using Perp_perm A B Perp D' B by blast
                    moreover
                    have "C'  B" 
                      using B Out C' A l6_3_1 by auto
                    moreover
                    have "Col B A C'" 
                      using Out_cases B Out C' A out_col by blast
                    moreover
                    have "Col B A B" 
                      by (simp add: col_trivial_3)
                    ultimately
                    show ?thesis 
                      using perp_col0 by blast
                  qed
                  moreover
                  have "B  P'" 
                    using B Out D' P' out_distinct by blast
                  moreover
                  have "Col B D' B" 
                    by (simp add: col_trivial_3)
                  moreover
                  have "Col B D' P'" 
                    using B Out D' P' out_col by auto
                  ultimately
                  show ?thesis 
                    using perp_col by blast
                qed
                hence "Per P' B C'" 
                  using perp_per_1 by blast
                moreover
                have "P'  B" 
                  using B Out D' P' out_diff2 by auto
                moreover
                have "C'  B" 
                  using P' C' B LtA A B C lta_distincts by blast
                moreover
                have "Per P'' C' B" 
                  using Per B C' P'' l8_2 by blast
                moreover
                have "P''  C'" 
                  using False Cong C' P'' P Q cong_diff_3 by blast
                moreover
                have "B  C'" 
                  using calculation(3) by auto
                ultimately
                show ?thesis 
                  by (simp add: l11_16)
              qed
              moreover
              have "Cong B P' C' P''" 
                by (meson Cong B P' P Q Cong C' P'' P Q
                    cong_inner_transitivity cong_symmetry)
              moreover
              have "Cong B C' C' B" 
                using cong_pseudo_reflexivity by blast
              moreover
              have "P'  C'" 
                using P' C' B LtA A B C lta_distincts by blast
              ultimately
              show ?thesis 
                using l11_49 by blast
            qed
            moreover
            have "C' Out B B" 
              using P' C' B LtA A B C 
                lta_distincts out_trivial by blast
            moreover
            have "C'  P'" 
              using P' C' B LtA A B C lta_distincts by blast
            hence "C' Out P' P'" 
              using out_trivial by auto
            moreover
            have "B Out A C'" 
              using Out_cases B Out C' A by auto
            moreover
            have "B Out P'' P''" 
              using B  P'' out_trivial by auto
            ultimately
            show ?thesis 
              using l11_10 by blast
          qed
          have "P'' InAngle Z' B Z"
          proof -
            have "P'' InAngle A B C" 
            proof -
              have "A B P'' LeA A B C" 
              proof -
                have "B C' P' LeA A B C" 
                  by (meson P' C' B LtA A B C 
                      lea_left_comm lta__lea)
                moreover
                have "B C' P' CongA A B P''" 
                  by (simp add: B C' P' CongA A B P'')
                moreover
                have "A B C CongA A B C" 
                  using Acute A B C acute_distincts
                    conga_refl by blast
                ultimately
                show ?thesis 
                  using l11_30 by blast
              qed
              moreover
              have "A B OS C P''" 
              proof -
                have "A B OS C D''" 
                proof -
                  have "A  B" 
                    using ¬ Col A B C not_col_distincts by blast
                  moreover
                  have "Col B C' A" 
                    by (simp add: B Out C' A out_col)
                  moreover
                  have "Col B C' B" 
                    by (simp add: col_trivial_3)
                  ultimately
                  show ?thesis 
                    using B C' OS C D'' col2_os__os by blast
                qed
                moreover
                have "A B OS D'' P''" 
                proof -
                  have "¬ Col B C' D''" 
                    using B C' OS C D'' col124__nos by blast
                  hence "¬ Col A B D''" 
                    using calculation col124__nos by blast
                  moreover
                  have "Col A B C'" 
                    using B Out C' A col_permutation_2 out_col by blast
                  moreover
                  have "C' Out D'' P''" 
                    using C' Out D'' P'' by auto
                  ultimately
                  show ?thesis using out_one_side_1 by blast
                qed
                ultimately
                show ?thesis 
                  using one_side_transitivity by blast
              qed
              ultimately
              show ?thesis 
                by (simp add: lea_in_angle)
            qed
            moreover
            have "B Out Z' A" 
              by (simp add: B Out A Z' l6_6)
            moreover
            have "B Out Z C" 
              using Out_cases B Out C Z by blast
            moreover
            have "B Out P'' P''" 
              using B  P'' out_trivial by auto
            ultimately
            show ?thesis 
              using l11_25 by blast
          qed
          hence H3: "Z'  B  Z  B  P''  B 
                 ( X. Bet Z' X Z  (X = B  B Out X P''))" 
            using InAngle_def by blast
          then consider T where "Bet Z' T Z" and "T = B  B Out T P''" 
            by blast

          have " T. Bet Z' T Z  (Bet B T P''  Bet B P'' T)" 
          proof - 
            have "T = B  (Bet B T P''  Bet B P'' T)" 
              using between_trivial2 by auto 
            moreover
            {
              assume "B Out T P''" 
              hence "Bet B T P''  Bet B P'' T" 
                using Out_def by blast
            }
            ultimately
            show ?thesis 
              by (metis Out_def 
                  thesis. (T. Bet Z' T Z; T = B  B Out T P''  thesis)  thesis 
                  between_trivial2)
          qed
          then obtain T where "Bet Z' T Z" and "Bet B T P''  Bet B P'' T" 
            by blast
          have "Z Z' Par C' P''" 
          proof -
            have "Coplanar B C' Z C'"     
              by (meson ncop_distincts)
            moreover
            have "Coplanar B C' Z P''" 
            proof -
              have "Coplanar B Z C C'" 
                using Out_cases B Out C Z out__coplanar by blast
              moreover
              have "Coplanar C' P'' D'' B" 
                using Out_cases C' Out D'' P'' out__coplanar by blast
              moreover
              have "Coplanar C' B D'' C" 
                using B C' OS C D'' ncoplanar_perm_7 os__coplanar by blast
              moreover have "¬ Col B C' D''" 
                using B C' OS C D'' one_side_not_col124 by auto
              moreover have "¬ Col B C' C" 
                using B C' OS C D'' col123__nos by blast
              moreover have "Coplanar B C' Z C" 
                using calculation(1) ncoplanar_perm_3 by blast
              moreover have "Coplanar C' B P'' D''" 
                using calculation(2) ncoplanar_perm_3 by blast
              moreover have "Coplanar C' C B D''" 
                using calculation(3) ncoplanar_perm_3 by blast
              ultimately
              show ?thesis 
                by (meson col_permutation_4 coplanar_perm_1 l9_30 ncop_distincts)
            qed
            moreover
            have "Coplanar B C' Z' C'" 
              by (meson ncop_distincts)
            moreover
            have "Coplanar B C' Z' P''" 
              using B Out A Z' B Out C' A l6_7 out__coplanar by blast
            moreover
            have "Z Z' Perp B C'" 
              by (metis Out_cases Perp_cases A B Perp Z Z' 
                  B C' Perp D'' C' B Out C' A out_col 
                  perp_col1 perp_not_eq_1)
            moreover
            have "C' P'' Perp B C'" 
              by (metis False Perp_perm B C' Perp D'' C' 
                  C' Out D'' P'' Cong C' P'' P Q cong_diff_3 
                  out_col perp_col1)
            ultimately
            show ?thesis 
              using l12_9 by blast
          qed
          {
            assume "T = Z"
            have "A B P'' CongA A B C" 
            proof -
              have "B Out A A" 
                using B Out C' A out_diff2 out_trivial by blast
              moreover
              have "B Out P'' C" 
                using B Out C Z Bet B T P''  Bet B P'' T 
                  Cong B Z B P'' T = Z between_cong 
                  cong_symmetry l6_6 by blast
              ultimately
              show ?thesis 
                by (meson l6_6 out2__conga)
            qed
            hence "P' C' B CongA A B C" 
              by (meson conga_trans B C' P' CongA A B P'' conga_left_comm)
            hence False 
              using P' C' B LtA A B C lta_not_conga by blast
          }
          hence "T  Z" 
            by auto
          have "T  Z'" 
            by (metis Col_cases A B OS C D' B C' P' CongA A B P''
                B Out A Z' B Out C' A B Out D' P'
                Bet B T P''  Bet B P'' T bet_col ncol_conga_ncol
                one_side_not_col123 one_side_not_col124 out_one_side)
          {
            assume "T = P''" 
            have "Acute B T Z" 
              by (metis B  P'' Cong B Z B P'' T = P''
                  T  Z cong__acute not_cong_3412)
            have "Obtuse B T Z" 
            proof -
              have "Acute B T Z'" 
              proof -
                have "Per B Z' T" 
                proof -
                  have "A B Perp T Z'" 
                  proof -
                    have "Z Z' Perp A B" 
                      by (meson Perp_perm A B Perp Z Z')
                    moreover
                    have "Col Z Z' T" 
                      by (simp add: Col_def Bet Z' T Z)
                    moreover
                    have "Col Z Z' Z'" 
                      using not_col_distincts by auto
                    ultimately
                    show ?thesis 
                      using T  Z' perp_col0 by blast
                  qed
                  moreover
                  have "Z'  B" 
                    by (simp add: H3)
                  moreover
                  have "Col A B B" 
                    by (simp add: col_trivial_2)
                  ultimately
                  show ?thesis 
                    by (meson Col A B Z' perp_col2 perp_per_1)
                qed
                hence "Per B Z' T  Obtuse B Z' T" 
                  by simp
                thus ?thesis 
                  by (metis acute_sym T  Z' H3 l11_43)
              qed
              moreover
              have "B T Z' SuppA B T Z" 
              proof-
                have "Z'  T" 
                  using T  Z' by auto
                moreover
                have "T  B" 
                  using B  P'' T = P'' by blast
                moreover
                have "T  Z" 
                  by (simp add: T  Z)
                moreover
                have "Bet Z' T Z" 
                  using Bet Z' T Z by auto
                ultimately
                show ?thesis 
                  by (simp add: bet__suppa suppa_left_comm)
              qed
              ultimately
              show ?thesis 
                using acute_suppa__obtuse by blast
            qed
            hence False 
              using B  P'' Cong B Z B P'' T = P'' 
                T  Z acute__not_obtuse cong__acute 
                not_cong_3412 by blast
          }
          hence "T  P''" 
            by blast
          {
            assume "Bet B P'' T"
            have "Obtuse Z T B" 
            proof -
              have "Acute B T Z'" 
              proof -
                have "Z'  B" 
                  using H3 by blast
                moreover
                have "Z'  T" 
                  using T  Z' by auto
                moreover
                have "Per B Z' T  Obtuse B Z' T" 
                  by (metis Per_perm l8_3 Bet Z' T Z Per B Z' Z 
                      bet_col bet_col1 bet_neq23__neq 
                      calculation(2) col_transitivity_1)
                ultimately
                show ?thesis 
                  by (meson acute_sym l11_43)
              qed
              moreover
              have "B  T Z' SuppA Z T B" 
              proof -
                have "Z'  T" 
                  using T  Z' by fastforce
                moreover
                have "T  B" 
                  using B  P'' Bet B P'' T bet_neq32__neq by blast
                ultimately
                show ?thesis 
                  by (simp add: T  Z Bet Z' T Z 
                      bet__suppa suppa_comm)
              qed
              ultimately
              show ?thesis 
                using acute_suppa__obtuse by blast
            qed
            have "B Z T LtA B T Z" 
            proof -
              have "Acute B Z T" 
              proof -
                have "T  B" 
                  using B  P'' Bet B P'' T between_identity by blast
                moreover
                have "Per B T Z  Obtuse B T Z" 
                  using Obtuse Z T B obtuse_sym by blast
                ultimately
                show ?thesis 
                  by (simp add:  T  Z Obtuse Z T B 
                      acute_sym l11_43_aux)
              qed
              moreover
              have "Obtuse B T Z" 
                by (simp add: Obtuse Z T B obtuse_sym)
              ultimately
              show ?thesis 
                using acute_obtuse__lta by blast
            qed
            moreover
            have "B T Z LtA B Z T" 
            proof -
              have "¬ Col Z B T" 
                by (metis NCol_perm Bet Z' T Z Col A B Z'
                    T  Z H3 ¬ Col A B Z bet_out_1 
                    col_trivial_2 l6_21 out_col)
              moreover
              have "B Z Lt B T" 
                by (metis nlt__le B  P'' Bet B P'' T 
                    Cong B Z B P'' T  P'' bet_out between_equality_2 
                    cong_reflexivity l5_6 l6_13_1 l6_6)
              ultimately
              show ?thesis 
                using col_permutation_3 l11_44_2 by blast
            qed
            ultimately
            have False 
              using not_and_lta by blast
          }
          hence "Bet B T P''" 
            using Bet B T P''  Bet B P'' T by auto
          have "C' P'' Par Z Z'" 
            by (simp add: Z Z' Par C' P'' par_symmetry)
          have "Col Z Z' T" 
            using Col_def Bet Z' T Z by blast
          have "¬ Col C' P'' T" 
            by (metis Col_cases H3 Bet B T P'' Col A B Z' 
                Col Z Z' T T = P''  False Z Z' Par C' P'' ¬ Col A B Z 
                bet_col col2__eq not_strict_par2)
          hence "C' P'' ParStrict Z Z'" 
            using par_not_col_strict C' P'' Par Z Z' 
              Col Z Z' T by blast
          have "Z Z' OS C' P''" 
            by (simp add: C' P'' ParStrict Z Z' pars__os3412)
          have "P'' Out T B" 
            by (simp add: Bet B T P'' T  P'' bet_out_1)
          have "B Out C' Z'" 
            using B Out A Z' B Out C' A l6_7 by blast
          hence "Bet B Z' C'  Bet B C' Z'" 
            using Out_def by force
          {
            assume "Bet B Z' C'"
            have "Z'  C'" 
              using Z Z' OS C' P'' col123__nos col_trivial_2 by blast
            hence "B Z' Lt B C'" 
              using bet__lt1213 by (simp add: Bet B Z' C')
          }
          moreover
          {
            assume "Bet B C' Z'"
            have "Z Z' TS B P''" 
            proof -
              have "¬ Col P'' Z Z'" 
                by (metis Col A B Z' Col Z Z' T 
                    P'' InAngle Z' B Z P'' Out T B ¬ Col A B Z 
                    ¬ Col C' P'' T col_permutation_1 colx 
                    inangle_distincts not_col_distincts out_col)
              moreover
              have "Col T Z Z'" 
                by (simp add: Col_def Bet Z' T Z)
              moreover
              have "Bet B T P''" 
                using Bet B T P'' by blast
              ultimately
              show ?thesis
                by (metis Col A B Z' P'' InAngle Z' B Z 
                    ¬ Col A B Z bet_col inangle_distincts l6_21 l9_18_R2
                    not_col_distincts not_col_permutation_2
                    not_col_permutation_5)
            qed
            moreover
            have "Z Z' OS B C'" 
              by (metis Bet B C' Z' Z Z' OS C' P'' 
                  bet_out between_symmetry col_trivial_2 l9_19_R2 
                  one_side_not_col123 one_side_symmetry)
            ultimately
            have False 
              using Z Z' OS C' P'' l9_9 one_side_transitivity by blast
            hence "B Z' Lt B C'" 
              by blast
          }
          thus ?thesis 
            using Bet B Z' C'  Bet B C' Z' calculation by fastforce
        qed
        ultimately
        show ?thesis 
          using  Cong B Z B P'' cong_lt_per2__lt_1 by blast
      qed
      hence "P Q Lt Z' Z" 
        by (meson Cong C' P'' P Q cong__nlt le3456_lt__lt nle__lt)
      ultimately
      show ?thesis 
        by blast
    qed
  }
  thus ?thesis 
    by (simp add: aristotle_s_axiom_def)
qed

theorem equiv_aristotle___greenberg: 
  shows  "aristotle_s_axiom  greenberg_s_axiom"
  using aristotle__greenberg greenberg__aristotle by blast

end
end