Theory Tarski_Neutral

(* IsageoCoq

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

Copyright (C) 2021  Roland Coghetto roland_coghetto (at) hotmail.com

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

imports
  Main

begin

section "Tarski's axiom system for neutral geometry"

subsection "Tarski's axiom system for neutral geometry: dimensionless"

locale Tarski_neutral_dimensionless =
  fixes Bet  :: "'p  'p  'p  bool"
  fixes Cong :: "'p  'p  'p  'p  bool"
  assumes cong_pseudo_reflexivity: " a b.
                                   Cong a b b a"
    and   cong_inner_transitivity: " a b p q r s.
                                    Cong a b p q 
                                    Cong a b r s
                                      
                                    Cong p q r s"
    and   cong_identity: " a b c.
                          Cong a b c c
                            
                          a = b"
    and   segment_construction: " a b c q.
                                 x. (Bet q a x  Cong a x b c)"
    and   five_segment: " a b c a' b' c'.
                         a  b 
                         Bet a b c 
                         Bet a' b' c'
                         Cong a b a' b' 
                         Cong b c b' c' 
                         Cong a d a' d' 
                         Cong b d b' d'
                           
                         Cong c d c' d'"
    and   between_identity: " a b.
                             Bet a b a
                               
                             a = b"
    and   inner_pasch: " a b c p q.
                        Bet a p c 
                        Bet b q c
                          
                        ( x. Bet p x b  Bet q x a)"
    and   lower_dim:  " a b c. (¬ Bet a b c  ¬ Bet b c a  ¬ Bet c a b)"

subsection "Tarski's axiom system for neutral geometry: 2D"

locale Tarski_2D = Tarski_neutral_dimensionless +
  assumes upper_dim: " a b c p q.
                      p  q 
                      Cong a p a q 
                      Cong b p b q 
                      Cong c p c q
                      
                      (Bet a b c  Bet b c a  Bet c a b)"

section "Definitions"
subsection "Tarski's axiom system for neutral geometry: dimensionless"
context Tarski_neutral_dimensionless
begin

subsubsection "Congruence"
definition OFSC ::
  "['p,'p,'p,'p,'p,'p,'p,'p]  bool"
  (‹_ _ _ _ OFSC _ _ _ _› [99,99,99,99,99,99,99,99] 50)
  where
    "A B C D OFSC A' B' C' D' 

  Bet A B C 
  Bet A' B' C' 
  Cong A B A' B' 
  Cong B C B' C' 
  Cong A D A' D' 
  Cong B D B' D'"

definition Cong3 ::
  "['p,'p,'p,'p,'p,'p]  bool"
  (‹_ _ _ Cong3 _ _ _› [99,99,99,99,99,99] 50)
  where
    "A B C Cong3 A' B' C' 

  Cong A B A' B' 
  Cong A C A' C' 
  Cong B C B' C'"

subsubsection "Betweenness"

definition Col ::
  "['p,'p,'p]  bool"
  (Col _ _ _› [99,99,99] 50)
  where
    "Col A B C 

  Bet A B C  Bet B C A  Bet C A B"

definition Bet4 ::
  "['p,'p,'p,'p]  bool"
  (Bet4 _ _ _ _› [99,99,99,99] 50)
  where
    "Bet4 A1 A2 A3 A4 

  Bet A1 A2 A3 
  Bet A2 A3 A4 
  Bet A1 A3 A4 
  Bet A1 A2 A4"

definition BetS ::
  "['p,'p,'p]  bool" (BetS _ _ _› [99,99,99] 50)
  where
    "BetS A B C 

  Bet A B C 
  A  B 
  B  C"

subsubsection "Collinearity"

definition FSC ::
  "['p,'p,'p,'p,'p,'p,'p,'p]  bool"
  (‹_ _ _ _ FSC _ _ _ _› [99,99,99,99,99,99,99,99] 50)
  where
    "A B C D FSC A' B' C' D' 

  Col A B C 
  A B C Cong3 A' B' C' 
  Cong A D A' D' 
  Cong B D B' D'"

subsubsection "Congruence and Betweenness"
definition IFSC ::
  "['p,'p,'p,'p,'p,'p,'p,'p]  bool"
  (‹_ _ _ _ IFSC _ _ _ _› [99,99,99,99,99,99,99,99] 50)
  where
    "A B C D IFSC A' B' C' D' 

  Bet A B C 
  Bet A' B' C' 
  Cong A C A' C' 
  Cong B C B' C' 
  Cong A D A' D' 
  Cong C D C' D'"

subsubsection "Between transivitity  LE"

definition Le ::
  "['p,'p,'p,'p]  bool" (‹_ _ Le _ _› [99,99,99,99] 50)
  where "A B Le C D 

   E. (Bet C E D  Cong A B C E)"


definition Lt ::
  "['p,'p,'p,'p]  bool" (‹_ _ Lt _ _› [99,99,99,99] 50)
  where "A B Lt C D 

  A B Le C D  ¬ Cong A B C D"

definition Ge ::
  "['p,'p,'p,'p]  bool" (‹_ _Ge _ _› [99,99,99,99] 50)
  where "A B Ge C D 

  C D Le A B"

definition Gt ::
  "['p,'p,'p,'p]  bool" (‹_ _ Gt _ _› [99,99,99,99] 50)
  where "A B Gt C D 

  C D Lt A B"

subsubsection "Out lines"

definition Out ::
  "['p,'p,'p]  bool" (‹_ Out _ _› [99,99,99] 50)
  where "P Out A B 

  A  P 
  B  P 
  (Bet P A B  Bet P B A)"

subsubsection "Midpoint"

definition Midpoint ::
  "['p,'p,'p]  bool" (‹_ Midpoint _ _› [99,99,99] 50)
  where "M Midpoint A B 

  Bet A M B 
  Cong A M M B"

subsubsection "Orthogonality"

definition Per ::
  "['p,'p,'p]  bool" (Per _ _ _› [99,99,99] 50)
  where "Per A B C 

   C'::'p. (B Midpoint C C'  Cong A C A C')"

definition PerpAt ::
  "['p,'p,'p,'p,'p]  bool" (‹_ PerpAt _ _ _ _ › [99,99,99,99,99] 50)
  where "X PerpAt A B C D 

  A  B 
  C  D 
  Col X A B 
  Col X C D 
  ( U V. ((Col U A B  Col V C D)  Per U X V))"

definition Perp ::
  "['p,'p,'p,'p]  bool" (‹_ _ Perp _ _› [99,99,99,99] 50)
  where "A B Perp C D 

   X::'p. X PerpAt A B C D"

subsubsection "Coplanar"

definition Coplanar ::
  "['p,'p,'p,'p]  bool" (Coplanar _ _ _ _› [99,99,99,99] 50)
  where "Coplanar A B C D 
   X. (Col A B X  Col C D X) 
            (Col A C X  Col B D X) 
            (Col A D X  Col B C X)"

definition TS ::
  "['p,'p,'p,'p]  bool" (‹_ _ TS _ _› [99,99,99,99] 50)
  where "A B TS P Q 
  ¬ Col P A B  ¬ Col Q A B  ( T::'p. Col T A B  Bet P T Q)"

definition ReflectL ::
  "['p,'p,'p,'p]  bool" (‹_ _ ReflectL _ _› [99,99,99,99] 50)
  where "P' P ReflectL A B 
  ( X. X Midpoint P P'  Col A B X)  (A B Perp P P'  P = P')"

definition Reflect ::
  "['p,'p,'p,'p]  bool" (‹_ _ Reflect _ _› [99,99,99,99] 50)
  where "P' P Reflect A B 
 (A  B  P' P ReflectL A B)  (A = B  A Midpoint P P')"

definition InAngle ::
  "['p,'p,'p,'p]  bool" (‹_ InAngle _ _ _› [99,99,99,99] 50)
  where "P InAngle A B C 
  A  B  C  B  P  B 
( X. Bet A X C  (X = B  B Out X P))"

definition ParStrict::
  "['p,'p,'p,'p]  bool" (‹_ _ ParStrict _ _› [99,99,99,99] 50)
  where "A B ParStrict C D   Coplanar A B C D  ¬ ( X. Col X A B  Col X C D)"

definition Par::
  "['p,'p,'p,'p]  bool" (‹_ _ Par _ _› [99,99,99,99] 50)
  where "A B Par C D 
  A B ParStrict C D  (A  B  C  D  Col A C D  Col B C D)"

definition Plg::
  "['p,'p,'p,'p]  bool" (Plg _ _ _ _› [99,99,99,99] 50)
  where "Plg  A B C D 
  (A  C  B  D)  ( M. M Midpoint A C  M Midpoint B D)"

definition ParallelogramStrict::
  "['p,'p,'p,'p]  bool" (ParallelogramStrict _ _ _ _› [99,99,99,99] 50)
  where "ParallelogramStrict A B A' B' 
  A A' TS B B'  A B Par A' B'  Cong A B A' B'"

definition ParallelogramFlat::
  "['p,'p,'p,'p]  bool" (ParallelogramFlat _ _ _ _› [99,99,99,99] 50)
  where "ParallelogramFlat A B A' B' 
  Col A B A'  Col A B B' 
  Cong A B A' B'  Cong A B' A' B 
  (A  A'  B  B')"

definition Parallelogram::
  "['p,'p,'p,'p]  bool" (Parallelogram _ _ _ _› [99,99,99,99] 50)
  where "Parallelogram A B A' B' 
  ParallelogramStrict A B A' B'  ParallelogramFlat A B A' B'"

definition Rhombus::
  "['p,'p,'p,'p]  bool" (Rhombus _ _ _ _› [99,99,99,99] 50)
  where "Rhombus A B C D  Plg A B C D  Cong A B B C"

definition Rectangle::
  "['p,'p,'p,'p]  bool" (Rectangle _ _ _ _› [99,99,99,99] 50)
  where "Rectangle A B C D  Plg A B C D  Cong A C B D"

definition Square::
  "['p,'p,'p,'p]  bool" (Square _ _ _ _› [99,99,99,99] 50)
  where "Square A B C D  Rectangle A B C D  Cong A B B C"

definition Lambert::
  "['p,'p,'p,'p]  bool" (Lambert _ _ _ _› [99,99,99,99] 50)
  where "Lambert A B C D 
  A  B  B  C  C  D 
 A  D  Per B A D  Per A D C  Per A B C  Coplanar A B C D"

subsubsection "Plane"

definition OS ::
  "['p,'p,'p,'p]  bool" (‹_ _ OS _ _› [99,99,99,99] 50)
  where "A B OS P Q 
 R::'p. A B TS P R  A B TS Q R"

definition TSP ::
  "['p,'p,'p,'p,'p]  bool" (‹_ _ _TSP _ _› [99,99,99,99,99] 50)
  where "A B C TSP P Q 
  (¬ Coplanar A B C P)  (¬ Coplanar A B C Q) 
( T. Coplanar A B C T  Bet P T Q)"

definition OSP ::
  "['p,'p,'p,'p,'p]  bool" (‹_ _ _ OSP _ _› [99,99,99,99,99] 50)
  where "A B C OSP P Q 
 R. ((A B C TSP P R)  (A B C TSP Q R))"

definition Saccheri::
  "['p,'p,'p,'p]  bool" (Saccheri _ _ _ _› [99,99,99,99] 50)
  where "Saccheri A B C D 
  Per B A D  Per A D C  Cong A B C D  A D OS B C"

subsubsection "Line reflexivity 2D"

definition ReflectLAt ::
  "['p,'p,'p,'p,'p]  bool" (‹_ ReflectLAt _ _ _ _› [99,99,99,99,99] 50)
  where "M ReflectLAt P' P A B 
  (M Midpoint P P'  Col A B M)  (A B Perp P P'  P = P')"

definition ReflectAt ::
  "['p,'p,'p,'p,'p]  bool" (‹_ ReflectAt _ _ _ _› [99,99,99,99,99] 50)
  where "M ReflectAt P' P A B 
(A  B  M ReflectLAt P' P A B)  (A = B  A = M  M Midpoint P P')"

subsubsection "Line reflexivity"

definition upper_dim_axiom ::
  "bool" (UpperDimAxiom [] 50)
  where
    "upper_dim_axiom 

   A B C P Q.
  P  Q 
  Cong A P A Q 
  Cong B P B Q 
  Cong C P C Q
    
  (Bet A B C  Bet B C A  Bet C A B)"

definition all_coplanar_axiom ::
  "bool" (AllCoplanarAxiom [] 50)
  where
    "AllCoplanarAxiom 

   A B C P Q.
  P  Q 
  Cong A P A Q 
  Cong B P B Q 
  Cong C P C Q
    
  (Bet A B C  Bet B C A  Bet C A B)"

subsubsection "Angles"

definition CongA ::
  "['p,'p,'p,'p,'p,'p]  bool" (‹_ _ _ CongA _ _ _› [99,99,99,99,99,99] 50)
  where "A B C CongA D E F 
  A  B  C  B  D  E  F  E 
( A' C' D' F'. Bet B A A'  Cong A A' E D 
  Bet B C C'  Cong C C' E F 
  Bet E D D'  Cong D D' B A 
  Bet E F F'  Cong F F' B C 
  Cong A' C' D' F')"

definition LeA ::
  "['p,'p,'p,'p,'p,'p]  bool" (‹_ _ _ LeA _ _ _› [99,99,99,99,99,99] 50)
  where "A B C LeA D E F 
 P. (P InAngle D E F  A B C CongA D E P)"

definition LtA ::
  "['p,'p,'p,'p,'p,'p]  bool" (‹_ _ _ LtA _ _ _› [99,99,99,99,99,99] 50)
  where "A B C LtA D E F  A B C LeA D E F  ¬ A B C CongA D E F"

definition GtA ::
  "['p,'p,'p,'p,'p,'p]  bool" (‹_ _ _ GtA _ _ _› [99,99,99,99,99,99] 50)
  where "A B C GtA D E F  D E F LtA A B C"

definition Acute ::
  "['p,'p,'p]  bool" (Acute _ _ _› [99,99,99] 50)
  where "Acute A B C 
 A' B' C'. (Per A' B' C'  A B C LtA A' B' C')"

definition Obtuse ::
  "['p,'p,'p]  bool" (Obtuse _ _ _› [99,99,99] 50)
  where "Obtuse A B C 
 A' B' C'. (Per A' B' C'  A' B' C' LtA A B C)"

definition OrthAt ::
  "['p,'p,'p,'p,'p,'p]  bool" (‹_ OrthAt _ _ _ _ _› [99,99,99,99,99,99] 50)
  where "X OrthAt A B C U V 
  ¬ Col A B C  U  V  Coplanar A B C X  Col U V X 
  ( P Q. (Coplanar A B C P  Col U V Q)  Per P X Q)"

definition Orth ::
  "['p,'p,'p,'p,'p]  bool" (‹_ _ _ Orth _ _› [99,99,99,99,99] 50)
  where "A B C Orth U V   X. X OrthAt A B C U V"

definition SuppA ::
  "['p,'p,'p,'p,'p,'p]  bool"
  (‹_ _ _ SuppA _ _ _ › [99,99,99,99,99,99] 50)
  where
    "A B C SuppA D E F 
  A  B  ( A'. Bet A B A'   D E F CongA C B A')"

subsubsection "Sum of angles"

definition SumA ::
  "['p,'p,'p,'p,'p,'p,'p,'p,'p]  bool" (‹_ _ _ _ _ _ SumA _ _ _› [99,99,99,99,99,99,99,99,99] 50)
  where
    "A B C D E F SumA G H I 

   J. (C B J CongA D E F  ¬ B C OS A J  Coplanar A B C J  A B J CongA G H I)"

definition TriSumA ::
  "['p,'p,'p,'p,'p,'p]  bool" (‹_ _ _ TriSumA _ _ _› [99,99,99,99,99,99] 50)
  where
    "A B C TriSumA D E F 

   G H I. (A B C B C A SumA G H I  G H I C A B SumA D E F)"

definition SAMS ::
  "['p,'p,'p,'p,'p,'p]  bool" (SAMS _ _ _ _ _ _› [99,99,99,99,99,99] 50)
  where
    "SAMS A B C D E F 

  (A  B 
  (E Out D F  ¬ Bet A B C)) 
  ( J. (C B J CongA D E F  ¬ (B C OS A J)  ¬ (A B TS C J)  Coplanar A B C J))"

subsubsection "Parallelism"

definition Inter ::
  "['p,'p,'p,'p,'p]  bool" (‹_ Inter _ _ _ _› [99,99,99,99,99] 50)
  where "X Inter A1 A2 B1 B2 

  B1  B2 
  ( P::'p. (Col P B1 B2  ¬ Col P A1 A2)) 
  Col A1 A2 X  Col B1 B2 X"

subsubsection "Perpendicularity"

definition Perp2 ::
  "['p,'p,'p,'p,'p]  bool" (‹_ Perp2 _ _ _ _› [99,99,99,99,99] 50)
  where
    "P Perp2 A B C D 

   X Y. (Col P X Y  X Y Perp A B  X Y Perp C D)"

subsubsection "Lentgh"

definition QCong::
  "(['p,'p]  bool)  bool" (QCong _› [99] 50)
  where
    "QCong l 

   A B. ( X Y. (Cong A B X Y  l X Y))"

definition TarskiLen::
  "['p,'p,(['p,'p]  bool)]  bool" (TarskiLen _ _ _› [99,99,99] 50)
  where
    "TarskiLen A B l 

  QCong l  l A B"

definition QCongNull ::
  "(['p,'p]  bool)  bool" (QCongNull _› [99] 50)
  where
    "QCongNull l 

  QCong l  ( A. l A A)"

subsubsection "Equivalence Class of Angles"

definition QCongA ::
  "(['p, 'p, 'p]  bool)  bool" (QCongA _› [99] 50)
  where
    "QCongA a 

   A B C. (A  B  C  B  ( X Y Z. A B C CongA X Y Z  a X Y Z))"

definition Ang ::
  "['p,'p,'p, (['p, 'p, 'p]  bool) ]  bool" (‹_ _ _ Ang _› [99,99,99,99] 50)
  where
    "A B C Ang a 

  QCongA a 
  a A B C"

definition QCongAAcute ::
  "(['p, 'p, 'p]  bool)  bool" (QCongAACute _› [99] 50)
  where
    "QCongAAcute a 

   A B C. (Acute A B C  ( X Y Z. (A B C CongA X Y Z  a X Y Z)))"

definition AngAcute ::
  "['p,'p,'p, (['p,'p,'p]  bool)]  bool" (‹_ _ _ AngAcute _› [99,99,99,99] 50)
  where
    "A B C AngAcute a 

  ((QCongAAcute a)  (a A B C))"

definition QCongANullAcute ::
  "(['p,'p,'p]  bool)  bool" (QCongANullAcute _› [99] 50)
  where
    "QCongANullAcute a 

   QCongAAcute a 
   ( A B C. (a A B C  B Out A C))"

definition QCongAnNull ::
  "(['p,'p,'p]  bool)  bool" (QCongAnNull _› [99] 50)
  where
    "QCongAnNull a 

  QCongA a 
  ( A B C. (a A B C  ¬ B Out A C))"

definition QCongAnFlat ::
  "(['p,'p,'p]  bool)  bool" (QCongAnFlat _› [99] 50)
  where
    "QCongAnFlat a 

  QCongA a 
  ( A B C. (a A B C  ¬ Bet A B C))"

definition IsNullAngaP ::
  "(['p,'p,'p]  bool)  bool" (IsNullAngaP _› [99] 50)
  where
    "IsNullAngaP a

  QCongAAcute a 
  ( A B C. (a A B C  B Out A C))"

definition QCongANull ::
  "(['p,'p,'p]  bool)  bool" (QCongANull _› [99] 50)
  where
    "QCongANull a 

  QCongA a 
  ( A B C. (a A B C  B Out A C))"

definition AngFlat ::
  "(['p, 'p, 'p]  bool)  bool" (AngFlat _› [99] 50)
  where
    "AngFlat a 

  QCongA a 
  ( A B C. (a A B C  Bet A B C))"

subsection "Parallel's definition Postulate"

definition tarski_s_parallel_postulate ::
  "bool"
  (TarskiSParallelPostulate)
  where
    "tarski_s_parallel_postulate 
 A B C D T. (Bet A D T  Bet B D C  A  D) 
( X Y. Bet A B X  Bet A C Y  Bet X T Y)"

definition euclid_5 ::
  "bool" (Euclid5)
  where
    "euclid_5 

   P Q R S T U.
  (BetS P T Q 
  BetS R T S 
  BetS Q U R 
  ¬ Col P Q S 
  Cong P T Q T 
  Cong R T S T)
    
  ( I. BetS S Q I  BetS P U I)"

definition euclid_s_parallel_postulate ::
  "bool" (EuclidSParallelPostulate)
  where
    "euclid_s_parallel_postulate 

   A B C D P Q R.
  (B C OS A D 
   SAMS A B C B C D 
   A B C B C D SumA P Q R 
   ¬ Bet P Q R)
    
  ( Y. B Out A Y  C Out D Y)"

definition playfair_s_postulate ::
  "bool"
  (PlayfairSPostulate)
  where
    "playfair_s_postulate 

   A1 A2 B1 B2 C1 C2 P.
  (A1 A2 Par B1 B2 
  Col P B1 B2 
  A1 A2 Par C1 C2 
  Col P C1 C2)
    
  (Col C1 B1 B2  Col C2 B1 B2)"

section "Propositions"

subsection "Congruence properties"

lemma cong_reflexivity:
  shows "Cong A B A B"
  using cong_inner_transitivity cong_pseudo_reflexivity by blast

lemma cong_symmetry:
  assumes "Cong A B C D"
  shows "Cong C D A B"
  using assms cong_inner_transitivity cong_reflexivity by blast

lemma cong_transitivity:
  assumes "Cong A B C D" and "Cong C D E F"
  shows "Cong A B E F"
  by (meson assms(1) assms(2) cong_inner_transitivity cong_pseudo_reflexivity)

lemma cong_left_commutativity:
  assumes "Cong A B C D"
  shows "Cong B A C D"
  using assms cong_inner_transitivity cong_pseudo_reflexivity by blast

lemma cong_right_commutativity:
  assumes "Cong A B C D"
  shows "Cong A B D C"
  using assms cong_left_commutativity cong_symmetry by blast

lemma cong_3421:
  assumes "Cong A B C D"
  shows "Cong C D B A"
  using assms cong_left_commutativity cong_symmetry by blast

lemma cong_4312:
  assumes "Cong A B C D"
  shows "Cong D C A B"
  using assms cong_left_commutativity cong_symmetry by blast

lemma cong_4321:
  assumes "Cong A B C D"
  shows "Cong D C B A"
  using assms cong_3421 cong_left_commutativity by blast

lemma cong_trivial_identity:
  shows "Cong A A B B"
  using cong_identity segment_construction by blast

lemma cong_reverse_identity:
  assumes "Cong A A C D"
  shows "C = D"
  using assms cong_3421 cong_identity by blast

lemma cong_commutativity:
  assumes "Cong A B C D"
  shows "Cong B A D C"
  using assms cong_3421 by blast

lemma not_cong_2134:
  assumes " ¬ Cong A B C D"
  shows "¬ Cong B A C D"
  using assms cong_left_commutativity by blast

lemma not_cong_1243:
  assumes "¬ Cong A B C D"
  shows "¬ Cong A B D C"
  using assms cong_right_commutativity by blast

lemma not_cong_2143:
  assumes "¬ Cong A B C D"
  shows "¬ Cong B A D C"
  using assms cong_commutativity by blast

lemma not_cong_3412:
  assumes "¬ Cong A B C D"
  shows "¬ Cong C D A B"
  using assms cong_symmetry by blast

lemma not_cong_4312:
  assumes "¬ Cong A B C D"
  shows "¬ Cong D C A B"
  using assms cong_3421 by blast

lemma not_cong_3421:
  assumes "¬ Cong A B C D"
  shows "¬ Cong C D B A"
  using assms cong_4312 by blast

lemma not_cong_4321:
  assumes "¬ Cong A B C D"
  shows "¬ Cong D C B A"
  using assms cong_4321 by blast

lemma five_segment_with_def:
  assumes "A B C D OFSC A' B' C' D'" and "A  B"
  shows "Cong C D C' D'"
  using assms(1) assms(2) OFSC_def five_segment by blast

lemma cong_diff:
  assumes "A  B" and "Cong A B C D"
  shows "C  D"
  using assms(1) assms(2) cong_identity by blast

lemma cong_diff_2:
  assumes "B  A" and "Cong A B C D"
  shows "C  D"
  using assms(1) assms(2) cong_identity by blast

lemma cong_diff_3:
  assumes "C  D" and "Cong A B C D"
  shows "A  B"
  using assms(1) assms(2) cong_reverse_identity by blast

lemma cong_diff_4:
  assumes "D  C" and "Cong A B C D"
  shows "A  B"
  using assms(1) assms(2) cong_reverse_identity by blast

lemma cong_3_sym:
  assumes "A B C Cong3 A' B' C'"
  shows "A' B' C' Cong3 A B C"
  using assms Cong3_def not_cong_3412 by blast

lemma cong_3_swap:
  assumes "A B C Cong3 A' B' C'"
  shows "B A C Cong3 B' A' C'"
  using assms Cong3_def cong_commutativity by blast

lemma cong_3_swap_2:
  assumes "A B C Cong3 A' B' C'"
  shows "A C B Cong3 A' C' B'"
  using assms Cong3_def cong_commutativity by blast

lemma cong3_transitivity:
  assumes "A0 B0 C0 Cong3 A1 B1 C1" and
    "A1 B1 C1 Cong3 A2 B2 C2"
  shows "A0 B0 C0 Cong3 A2 B2 C2"
  by (meson assms(1) assms(2) Cong3_def cong_inner_transitivity not_cong_3412)

lemma eq_dec_points:
  shows "A = B  ¬ A = B"
  by simp

lemma distinct:
  assumes "P  Q"
  shows "R  P  R  Q"
  using assms by simp

lemma l2_11:
  assumes "Bet A B C" and
    "Bet A' B' C'" and
    "Cong A B A' B'" and
    "Cong B C B' C'"
  shows "Cong A C A' C'"
  by (smt assms(1) assms(2) assms(3) assms(4) cong_right_commutativity cong_symmetry cong_trivial_identity five_segment)

lemma bet_cong3:
  assumes "Bet A B C" and
    "Cong A B A' B'"
  shows " C'. A B C Cong3 A' B' C'"
  by (meson assms(1) assms(2) Cong3_def l2_11 not_cong_3412 segment_construction)

lemma construction_uniqueness:
  assumes "Q  A" and
    "Bet Q A X" and
    "Cong A X B C" and
    "Bet Q A Y" and
    "Cong A Y B C"
  shows "X = Y"
  by (meson assms(1) assms(2) assms(3) assms(4) assms(5) cong_identity cong_inner_transitivity cong_reflexivity five_segment)

lemma Cong_cases:
  assumes "Cong A B C D  Cong A B D C  Cong B A C D  Cong B A D C  Cong C D A B  Cong C D B A  Cong D C A B  Cong D C B A"
  shows "Cong A B C D"
  using assms not_cong_3421 not_cong_4321 by blast

lemma Cong_perm :
  assumes "Cong A B C D"
  shows "Cong A B C D  Cong A B D C  Cong B A C D  Cong B A D C  Cong C D A B  Cong C D B A  Cong D C A B  Cong D C B A"
  using assms not_cong_1243 not_cong_3412 by blast

subsection "Betweeness properties"

lemma bet_col:
  assumes "Bet A B C"
  shows "Col A B C"
  by (simp add: assms Col_def)

lemma between_trivial:
  shows "Bet A B B"
  using cong_identity segment_construction by blast

lemma between_symmetry:
  assumes "Bet A B C"
  shows "Bet C B A"
  using assms between_identity between_trivial inner_pasch by blast

lemma Bet_cases:
  assumes "Bet A B C  Bet C B A"
  shows "Bet A B C"
  using assms between_symmetry by blast

lemma Bet_perm:
  assumes "Bet A B C"
  shows "Bet A B C  Bet C B A"
  using assms Bet_cases by blast

lemma between_trivial2:
  shows "Bet A A B"
  using Bet_perm between_trivial by blast

lemma between_equality:
  assumes "Bet A B C" and "Bet B A C"
  shows "A = B"
  using assms(1) assms(2) between_identity inner_pasch by blast

lemma between_equality_2:
  assumes "Bet A B C" and
    "Bet A C B"
  shows "B = C"
  using assms(1) assms(2) between_equality between_symmetry by blast

lemma between_exchange3:
  assumes "Bet A B C" and
    "Bet A C D"
  shows "Bet B C D"
  by (metis Bet_perm assms(1) assms(2) between_identity inner_pasch)

lemma bet_neq12__neq:
  assumes "Bet A B C" and
    "A  B"
  shows "A  C"
  using assms(1) assms(2) between_identity by blast

lemma bet_neq21__neq:
  assumes "Bet A B C" and
    "B  A"
  shows "A  C"
  using assms(1) assms(2) between_identity by blast

lemma bet_neq23__neq:
  assumes "Bet A B C" and
    "B  C"
  shows "A  C"
  using assms(1) assms(2) between_identity by blast

lemma bet_neq32__neq:
  assumes "Bet A B C" and
    "C  B"
  shows "A  C"
  using assms(1) assms(2) between_identity by blast

lemma not_bet_distincts:
  assumes "¬ Bet A B C"
  shows "A  B  B  C"
  using assms between_trivial between_trivial2 by blast

lemma between_inner_transitivity:
  assumes "Bet A B D" and
    "Bet B C D"
  shows "Bet A B C"
  using assms(1) assms(2) Bet_perm between_exchange3 by blast

lemma outer_transitivity_between2:
  assumes "Bet A B C" and
    "Bet B C D" and
    "B  C"
  shows "Bet A C D"
proof -
  obtain X where "Bet A C X  Cong C X C D"
    using segment_construction by blast
  thus ?thesis
    using assms(1) assms(2) assms(3) between_exchange3 cong_inner_transitivity construction_uniqueness by blast
qed

lemma between_exchange2:
  assumes "Bet A B D" and
    "Bet B C D"
  shows "Bet A C D"
  using assms(1) assms(2) between_inner_transitivity outer_transitivity_between2 by blast

lemma outer_transitivity_between:
  assumes "Bet A B C" and
    "Bet B C D" and
    "B  C"
  shows "Bet A B D"
  using assms(1) assms(2) assms(3) between_symmetry outer_transitivity_between2 by blast

lemma between_exchange4:
  assumes "Bet A B C" and
    "Bet A C D"
  shows "Bet A B D"
  using assms(1) assms(2) between_exchange2 between_symmetry by blast

lemma l3_9_4:
  assumes "Bet4 A1 A2 A3 A4"
  shows "Bet4 A4 A3 A2 A1"
  using assms Bet4_def Bet_cases by blast

lemma l3_17:
  assumes "Bet A B C" and
    "Bet A' B' C" and
    "Bet A P A'"
  shows "( Q. Bet P Q C  Bet B Q B')"
proof -
  obtain X where "Bet B' X A  Bet P X C"
    using Bet_perm assms(2) assms(3) inner_pasch by blast
  moreover
  then obtain Y where "Bet X Y C  Bet B Y B'"
    using Bet_perm assms(1) inner_pasch by blast
  ultimately show ?thesis
    using between_exchange2 by blast
qed

lemma lower_dim_ex:
  " A B C. ¬ (Bet A B C  Bet B C A  Bet C A B)"
  using lower_dim by auto

lemma two_distinct_points:
  " X::'p.  Y::'p. X  Y"
  using lower_dim_ex not_bet_distincts by blast

lemma point_construction_different:
  " C. Bet A B C  B  C"
  using Tarski_neutral_dimensionless.two_distinct_points Tarski_neutral_dimensionless_axioms cong_reverse_identity segment_construction by blast

lemma another_point:
  " B::'p. A  B"
  using point_construction_different by blast

lemma Cong_stability:
  assumes "¬ ¬ Cong A B C D"
  shows "Cong A B C D"
  using assms by simp

lemma l2_11_b:
  assumes "Bet A B C" and
    "Bet A' B' C'" and
    "Cong A B A' B'" and
    "Cong B C B' C'"
  shows "Cong A C A' C'"
  using assms(1) assms(2) assms(3) assms(4) l2_11 by auto

lemma cong_dec_eq_dec_b:
  assumes "¬ A  B"
  shows "A = B"
  using assms(1) by simp

lemma BetSEq:
  assumes "BetS A B C"
  shows "Bet A B C  A  B  A  C  B  C"
  using assms BetS_def between_identity by auto

subsection "Collinearity"

subsubsection "Collinearity and betweenness"

lemma l4_2:
  assumes "A B C D IFSC A' B' C' D'"
  shows "Cong B D B' D'"
proof cases
  assume "A = C"
  thus ?thesis
    by (metis IFSC_def Tarski_neutral_dimensionless.between_identity Tarski_neutral_dimensionless_axioms assms cong_diff_3)
next
  assume H1: "A  C"
  have H2: "Bet A B C  Bet A' B' C' 
            Cong A C A' C'  Cong B C B' C' 
            Cong A D A' D'  Cong C D C' D'"
    using IFSC_def assms by auto
  obtain E where P1: "Bet A C E  Cong C E A C"
    using segment_construction by blast
  have P1A: "Bet A C E"
    using P1 by simp
  have P1B: "Cong C E A C"
    using P1 by simp
  obtain E' where P2: "Bet A' C' E'  Cong C' E' C E"
    using segment_construction by blast
  have P2A: "Bet A' C' E'"
    using P2 by simp
  have P2B: "Cong C' E' C E"
    using P2 by simp
  then have "Cong C E C' E'"
    using not_cong_3412 by blast
  then have "Cong E D E' D'"
    using H1 H2 P1 P2 five_segment by blast
  thus ?thesis
    by (smt H1 H2 P1A P1B P2A P2B Tarski_neutral_dimensionless.cong_commutativity Tarski_neutral_dimensionless.cong_diff_3 Tarski_neutral_dimensionless.cong_symmetry Tarski_neutral_dimensionless_axioms between_inner_transitivity between_symmetry five_segment)
qed

lemma l4_3:
  assumes "Bet A B C" and
    "Bet A' B' C'" and
    "Cong A C A' C'"
    and "Cong B C B' C'"
  shows "Cong A B A' B'"
proof -
  have "A B C A IFSC A' B' C' A'"
    using IFSC_def assms(1) assms(2) assms(3) assms(4) cong_trivial_identity not_cong_2143 by blast
  thus ?thesis
    using l4_2 not_cong_2143 by blast
qed


lemma l4_3_1:
  assumes "Bet A B C" and
    "Bet A' B' C'" and
    "Cong A B A' B'" and
    "Cong A C A' C'"
  shows "Cong B C B' C'"
  by (meson assms(1) assms(2) assms(3) assms(4) between_symmetry cong_4321 l4_3)

lemma l4_5:
  assumes "Bet A B C" and
    "Cong A C A' C'"
  shows  " B'. (Bet A' B' C'  A B C Cong3 A' B' C')"
proof -
  obtain X' where P1: "Bet C' A' X'  A'  X'"
    using point_construction_different by auto
  obtain B' where P2: "Bet  X' A' B'  Cong A' B' A B"
    using segment_construction by blast
  obtain C'' where P3: "Bet X' B' C''  Cong B' C'' B C"
    using segment_construction by blast
  then have P4: "Bet A' B' C''"
    using P2 between_exchange3 by blast
  then have "C'' = C'"
    by (smt P1 P2 P3 assms(1) assms(2) between_exchange4 between_symmetry cong_symmetry construction_uniqueness l2_11_b)
  then show ?thesis
    by (smt Cong3_def P1 P2 P3 Tarski_neutral_dimensionless.construction_uniqueness Tarski_neutral_dimensionless_axioms P4 assms(1) assms(2) between_exchange4 between_symmetry cong_commutativity cong_symmetry cong_trivial_identity five_segment not_bet_distincts)
qed

lemma l4_6:
  assumes "Bet A B C" and
    "A B C Cong3 A' B' C'"
  shows "Bet A' B' C'"
proof -
  obtain x where P1: "Bet A' x C'  A B C Cong3 A' x C'"
    using Cong3_def assms(1) assms(2) l4_5 by blast
  then have "A' x C' Cong3 A' B' C'"
    using assms(2) cong3_transitivity cong_3_sym by blast
  then have "A' x C' x IFSC A' x C' B'"
    by (meson Cong3_def Cong_perm IFSC_def P1 cong_reflexivity)
  then have "Cong x x x B'"
    using l4_2 by auto
  then show ?thesis
    using P1 cong_reverse_identity by blast
qed

lemma cong3_bet_eq:
  assumes "Bet A B C" and
    "A B C Cong3 A X C"
  shows "X = B"
proof -
  have "A B C B IFSC A B C X"
    by (meson Cong3_def Cong_perm IFSC_def assms(1) assms(2) cong_reflexivity)
  then show ?thesis
    using cong_reverse_identity l4_2 by blast
qed

subsubsection "Collinearity"

lemma col_permutation_1:
  assumes "Col A B C"
  shows "Col B C A"
  using assms(1) Col_def by blast

lemma col_permutation_2:
  assumes "Col A B C"
  shows "Col C A B"
  using assms(1) col_permutation_1 by blast

lemma col_permutation_3:
  assumes "Col A B C"
  shows "Col C B A"
  using assms(1) Bet_cases Col_def by auto

lemma col_permutation_4:
  assumes "Col A B C"
  shows "Col B A C"
  using assms(1) Bet_perm Col_def by blast

lemma col_permutation_5:
  assumes "Col A B C"
  shows "Col A C B"
  using assms(1) col_permutation_1 col_permutation_3 by blast

lemma not_col_permutation_1:
  assumes "¬ Col A B C"
  shows "¬ Col B C A"
  using assms col_permutation_2 by blast

lemma not_col_permutation_2:
  assumes "~ Col A B C"
  shows  "~ Col C A B"
  using assms col_permutation_1 by blast

lemma not_col_permutation_3:
  assumes "¬ Col A B C"
  shows "¬ Col C B A"
  using assms col_permutation_3 by blast

lemma not_col_permutation_4:
  assumes "¬ Col A B C"
  shows "¬ Col B A C"
  using assms col_permutation_4 by blast

lemma not_col_permutation_5:
  assumes "¬ Col A B C"
  shows "¬ Col A C B"
  using assms col_permutation_5 by blast

lemma Col_cases:
  assumes "Col A B C  Col A C B  Col B A C  Col B C A  Col C A B  Col C B A"
  shows "Col A B C"
  using assms not_col_permutation_4 not_col_permutation_5 by blast

lemma Col_perm:
  assumes "Col A B C"
  shows "Col A B C  Col A C B  Col B A C  Col B C A  Col C A B  Col C B A"
  using Col_cases assms by blast

lemma col_trivial_1:
  "Col A A B"
  using bet_col not_bet_distincts by blast

lemma col_trivial_2:
  "Col A B B"
  by (simp add: Col_def between_trivial2)

lemma col_trivial_3:
  "Col A B A"
  by (simp add: Col_def between_trivial2)

lemma l4_13:
  assumes "Col A B C" and
    "A B C Cong3 A' B' C'"
  shows "Col A' B' C'"
  by (metis Tarski_neutral_dimensionless.Col_def Tarski_neutral_dimensionless.cong_3_swap Tarski_neutral_dimensionless.cong_3_swap_2 Tarski_neutral_dimensionless_axioms assms(1) assms(2) l4_6)

lemma l4_14R1:
  assumes "Bet A B C" and
    "Cong A B A' B'"
  shows " C'. A B C Cong3 A' B' C'"
  by (simp add: assms(1) assms(2) bet_cong3)

lemma l4_14R2:
  assumes "Bet B C A" and
    "Cong A B A' B'"
  shows " C'. A B C Cong3 A' B' C'"
  by (meson assms(1) assms(2) between_symmetry cong_3_swap_2 l4_5)

lemma l4_14R3:
  assumes "Bet C A B" and
    "Cong A B A' B'"
  shows " C'. A B C Cong3 A' B' C'"
  by (meson assms(1) assms(2) between_symmetry cong_3_swap l4_14R1 not_cong_2143)

lemma l4_14:
  assumes "Col A B C" and
    "Cong A B A' B'"
  shows " C'. A B C Cong3 A' B' C'"
  using Col_def assms(1) assms(2) l4_14R1 l4_14R2 l4_14R3 by blast

lemma l4_16R1:
  assumes "A B C D FSC A' B' C' D'" and
    "A  B" and
    "Bet A B C"
  shows "Cong C D C' D'"
proof -
  have "A B C Cong3 A' B' C'"
    using FSC_def assms(1) by blast
  then have "Bet A' B' C'"
    using assms(3) l4_6 by blast
  then have "A B C D OFSC A' B' C' D'"
    by (meson Cong3_def FSC_def OFSC_def assms(1) cong_3_sym l4_6)
  thus ?thesis
    using assms(2) five_segment_with_def by blast
qed

lemma l4_16R2:
  assumes "A B C D FSC A' B' C' D'"
    and "Bet B C A"
  shows "Cong C D C' D'"
proof -
  have "A B C Cong3 A' B' C'"
    using FSC_def assms(1) by blast
  then have "Bet B' C' A'"
    using Bet_perm assms(2) cong_3_swap_2 l4_6 by blast
  then have "B C A D IFSC B' C' A' D'"
    by (meson Cong3_def FSC_def IFSC_def assms(1) assms(2) not_cong_2143)
  then show ?thesis
    using l4_2 by auto
qed

lemma l4_16R3:
  assumes "A B C D FSC A' B' C' D'" and "A  B"
    and "Bet C A B"
  shows "Cong C D C' D'"
proof -
  have "A B C Cong3 A' B' C'"
    using FSC_def assms(1) by blast
  then have "Bet C' A' B'"
    using assms(3) between_symmetry cong_3_swap l4_6 by blast
  thus ?thesis
    by (smt Cong3_def FSC_def assms(1) assms(2) assms(3) between_symmetry cong_commutativity five_segment)
qed

lemma l4_16:
  assumes "A B C D FSC A' B' C' D'" and
    "A  B"
  shows "Cong C D C' D'"
  by (meson Col_def FSC_def assms(1) assms(2) l4_16R1 l4_16R2 l4_16R3)

lemma l4_17:
  assumes "A  B" and
    "Col A B C" and
    "Cong A P A Q" and
    "Cong B P B Q"
  shows "Cong C P C Q"
proof -
  {
    assume "¬ Bet B C A"
    then have "p pa. Bet p pa C  Cong pa P pa Q  Cong p P p Q  p  pa"
      using Col_def assms(1) assms(2) assms(3) assms(4) between_symmetry by blast
    then have ?thesis
      using cong_reflexivity five_segment by blast
  }
  then show ?thesis
    by (meson IFSC_def assms(3) assms(4) cong_reflexivity l4_2)
qed


lemma l4_18:
  assumes "A  B" and
    "Col A B C" and
    "Cong A C A C'" and
    "Cong B C B C'"
  shows "C = C'"
  using assms(1) assms(2) assms(3) assms(4) cong_diff_3 l4_17 by blast

lemma l4_19:
  assumes "Bet A C B" and
    "Cong A C A C'" and
    "Cong B C B C'"
  shows "C = C'"
  by (metis Col_def assms(1) assms(2) assms(3) between_equality between_trivial cong_identity l4_18 not_cong_3421)

lemma not_col_distincts:
  assumes "¬ Col A B C"
  shows "¬ Col A B C  A  B  B  C  A  C"
  using Col_def assms between_trivial by blast

lemma NCol_cases:
  assumes "¬ Col A B C  ¬ Col A C B  ¬ Col B A C  ¬ Col B C A  ¬ Col C A B  ¬ Col C B A"
  shows  "¬ Col A B C"
  using assms not_col_permutation_2 not_col_permutation_3 by blast

lemma NCol_perm:
  assumes "¬ Col A B C"
  shows "¬ Col A B C  ~ Col A C B  ~ Col B A C  ~ Col B C A  ~ Col C A B  ~ Col C B A"
  using NCol_cases assms by blast

lemma col_cong_3_cong_3_eq:
  assumes "A  B"
    and "Col A B C"
    and "A B C Cong3 A' B' C1"
    and  "A B C Cong3 A' B' C2"
  shows  "C1 = C2"
  by (metis Tarski_neutral_dimensionless.Cong3_def Tarski_neutral_dimensionless.cong_diff Tarski_neutral_dimensionless.l4_18 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) assms(4) cong_inner_transitivity l4_13)

subsection "Between transitivity le"

lemma l5_1:
  assumes "A  B" and
    "Bet A B C" and
    "Bet A B D"
  shows "Bet A C D  Bet A D C"
proof -
  obtain C' where P1: "Bet A D C'  Cong D C' C D"
    using segment_construction by blast
  obtain D' where P2: "Bet A C D'  Cong C D' C D"
    using segment_construction by blast
  obtain B' where P3: "Bet A C' B'  Cong C' B' C B"
    using segment_construction by blast
  obtain B'' where P4: "Bet A D' B''  Cong D' B'' D B"
    using segment_construction by blast
  then have P5: "Cong B C' B'' C"
    by (smt P1 P2 assms(3) between_exchange3 between_symmetry cong_4312 cong_inner_transitivity l2_11_b)
  then have "Cong B B' B'' B"
    by (meson Bet_cases P1 P2 P3 P4 assms(2) assms(3) between_exchange4 between_inner_transitivity l2_11_b)
  then have P6: "B'' = B'"
    by (meson P1 P2 P3 P4 assms(1) assms(2) assms(3) between_exchange4 cong_inner_transitivity construction_uniqueness not_cong_2134)
  have "Bet B C D'"
    using P2 assms(2) between_exchange3 by blast
  then have "B C D' C' FSC B' C' D C"
    by (smt Cong3_def FSC_def P1 P2 P3 P5 P6 bet_col between_exchange3 between_symmetry cong_3421 cong_pseudo_reflexivity cong_transitivity l2_11_b)
  then have P8: "Cong D' C' D C"
    using P3 P4 P6 cong_identity l4_16 by blast
  obtain E where P9: "Bet C E C'  Bet D E D'"
    using P1 P2 between_trivial2 l3_17 by blast
  then have P10: "D E D' C IFSC D E D' C'"
    by (smt IFSC_def P1 P2 P8 Tarski_neutral_dimensionless.cong_reflexivity Tarski_neutral_dimensionless_axioms cong_3421 cong_inner_transitivity)
  then have "Cong E C E C'"
    using l4_2 by auto
  have P11: "C E C' D IFSC C E C' D'"
    by (smt IFSC_def P1 P2 Tarski_neutral_dimensionless.cong_reflexivity Tarski_neutral_dimensionless_axioms P8 P9 cong_3421 cong_inner_transitivity)
  then have "Cong E D E D'"
    using l4_2 by auto
  obtain P where "Bet C' C P  Cong C P C D'"
    using segment_construction by blast
  obtain R where "Bet D' C R  Cong C R C E"
    using segment_construction by blast
  obtain Q where "Bet P R Q  Cong R Q R P"
    using segment_construction by blast
  have "D' C R P FSC P C E D'"
    by (meson Bet_perm Cong3_def FSC_def Bet C E C'  Bet D E D' Bet C' C P  Cong C P C D' Bet D' C R  Cong C R C E bet_col between_exchange3 cong_pseudo_reflexivity l2_11_b not_cong_4321)
  have "Cong R P E D'"
    by (metis Cong_cases D' C R P FSC P C E D' Bet C' C P  Cong C P C D' Bet D' C R  Cong C R C E cong_diff_2 l4_16)
  have "Cong R Q E D"
    by (metis Cong_cases Cong E D E D' Cong R P E D' Bet P R Q  Cong R Q R P cong_transitivity)
  have "D' E D C FSC P R Q C"
    by (meson Bet_perm Cong3_def FSC_def Cong R P E D' Cong R Q E D Bet C E C'  Bet D E D' Bet C' C P  Cong C P C D' Bet D' C R  Cong C R C E Bet P R Q  Cong R Q R P bet_col l2_11_b not_cong_2143 not_cong_4321)
  have "Cong D C Q C"
    using D' E D C FSC P R Q C Cong E D E D' Bet C E C'  Bet D E D' cong_identity l4_16 l4_16R2 by blast
  have "Cong C P C Q"
    using P2 Cong D C Q C Bet C' C P  Cong C P C D' cong_right_commutativity cong_transitivity by blast
  have "Bet A C D  Bet A D C"
  proof cases
    assume "R = C"
    then show ?thesis
      by (metis P1 Cong E C E C' Bet D' C R  Cong C R C E cong_diff_4)
  next
    assume "R  C"
    {
      have "Cong D' P D' Q"
      proof -

        have "Col R C D'"
          by (simp add: Bet D' C R  Cong C R C E bet_col between_symmetry)
        have "Cong R P R Q"
          by (metis Tarski_neutral_dimensionless.Cong_cases Tarski_neutral_dimensionless_axioms Bet P R Q  Cong R Q R P)
        have "Cong C P C Q"
          by (simp add: Cong C P C Q)
        then show ?thesis
          using Col R C D' Cong R P R Q R  C l4_17 by blast
      qed
      then have "Cong B P B Q" using  Cong C P C Q Bet B C D' cong_diff_4
        by (metis Col_def Bet C' C P  Cong C P C D' cong_reflexivity l4_17 not_cong_3412)
      have "Cong B' P B' Q"
        by (metis P2 P4 B'' = B' Cong C P C Q Cong D' P D' Q Bet C' C P  Cong C P C D' between_exchange3 cong_diff_4 cong_identity cong_reflexivity five_segment)
      have "Cong C' P C' Q"
      proof -
        have "Bet B C' B'"
          using P1 P3 assms(3) between_exchange3 between_exchange4 by blast
        then show ?thesis
          by (metis Col_def Cong B P B Q Cong B' P B' Q between_equality l4_17 not_bet_distincts)
      qed
      have "Cong P P P Q"
        by (metis Tarski_neutral_dimensionless.cong_diff_2 Tarski_neutral_dimensionless_axioms Cong C P C Q Cong C' P C' Q R  C Bet C E C'  Bet D E D' Bet C' C P  Cong C P C D' Bet D' C R  Cong C R C E bet_col bet_neq12__neq l4_17)
      thus ?thesis
        by (metis P2 Cong R P E D' Cong R Q E D Bet P R Q  Cong R Q R P bet_neq12__neq cong_diff_4)
    }
    then have "R  C  Bet A C D  Bet A D C" by blast
  qed
  thus ?thesis
    by simp
qed

lemma l5_2:
  assumes "A  B" and
    "Bet A B C" and
    "Bet A B D"
  shows  "Bet B C D  Bet B D C"
  using assms(1) assms(2) assms(3) between_exchange3 l5_1 by blast

lemma segment_construction_2:
  assumes  "A  Q"
  shows " X. ((Bet Q A X  Bet Q X A)  Cong Q X B C)"
proof -
  obtain A' where P1: "Bet A Q A'  Cong Q A' A Q"
    using segment_construction by blast
  obtain X where P2: "Bet A' Q X  Cong Q X B C"
    using segment_construction by blast
  then show ?thesis
    by (metis P1 Tarski_neutral_dimensionless.cong_diff_4 Tarski_neutral_dimensionless_axioms between_symmetry l5_2)
qed

lemma l5_3:
  assumes "Bet A B D" and
    "Bet A C D"
  shows "Bet A B C  Bet A C B"
  by (metis Bet_perm assms(1) assms(2) between_inner_transitivity l5_2 point_construction_different)

lemma bet3__bet:
  assumes "Bet A B E" and
    "Bet A D E" and
    "Bet B C D"
  shows "Bet A C E"
  by (meson assms(1) assms(2) assms(3) between_exchange2 between_symmetry l5_3)

lemma le_bet:
  assumes "C D Le A B"
  shows " X. (Bet A X B  Cong A X C D)"
  by (meson Le_def assms cong_symmetry)

lemma l5_5_1:
  assumes "A B Le C D"
  shows " X. (Bet A B X  Cong A X C D)"
proof -
  obtain P where P1: "Bet C P D  Cong A B C P"
    using Le_def assms by blast
  obtain X where P2: "Bet A B X  Cong B X P D"
    using segment_construction by blast
  then have "Cong A X C D"
    using P1 l2_11_b by blast
  then show ?thesis
    using P2 by blast
qed

lemma l5_5_2:
  assumes " X. (Bet A B X  Cong A X C D)"
  shows "A B Le C D"
proof -
  obtain P where P1: "Bet A B P  Cong A P C D"
    using assms by blast
  obtain B' where P2: "Bet C B' D  A B P Cong3 C B' D"
    using P1 l4_5 by blast
  then show ?thesis
    using Cong3_def Le_def by blast
qed

lemma l5_6:
  assumes "A B Le C D" and
    "Cong A B A' B'" and
    "Cong C D C' D'"
  shows "A' B' Le C' D'"
  by (meson Cong3_def Le_def assms(1) assms(2) assms(3) cong_inner_transitivity l4_5)

lemma le_reflexivity:
  shows "A B Le A B"
  using between_trivial cong_reflexivity l5_5_2 by blast

lemma le_transitivity:
  assumes "A B Le C D" and
    "C D Le E F"
  shows "A B Le E F"
  by (meson assms(1) assms(2) between_exchange4 cong_reflexivity l5_5_1 l5_5_2 l5_6 le_bet)

lemma between_cong:
  assumes "Bet A C B" and
    "Cong A C A B"
  shows "C = B"
  by (smt assms(1) assms(2) between_trivial cong_inner_transitivity five_segment l4_19 l4_3_1)

lemma cong3_symmetry:
  assumes "A B C Cong3 A' B' C'"
  shows "A' B' C' Cong3 A B C"
  by (simp add: assms cong_3_sym)

lemma between_cong_2:
  assumes "Bet A D B" and
    "Bet A E B"
    and "Cong A D A E"
  shows "D = E" using l5_3
  by (smt Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) cong_diff cong_inner_transitivity l4_3_1)

lemma between_cong_3:
  assumes "A  B"
    and "Bet A B D"
    and "Bet A B E"
    and "Cong B D B E"
  shows "D = E"
  by (meson assms(1) assms(2) assms(3) assms(4) cong_reflexivity construction_uniqueness)

lemma le_anti_symmetry:
  assumes "A B Le C D" and
    "C D Le A B"
  shows "Cong A B C D"
  by (smt Le_def Tarski_neutral_dimensionless.between_exchange4 Tarski_neutral_dimensionless_axioms assms(1) assms(2) bet_neq21__neq between_cong between_exchange3 cong_transitivity l5_5_1 not_cong_3421)

lemma cong_dec:
  shows "Cong A B C D  ¬ Cong A B C D"
  by simp

lemma bet_dec:
  shows "Bet A B C   ¬ Bet A B C"
  by simp

lemma col_dec:
  shows "Col A B C  ¬ Col A B C"
  by simp

lemma le_trivial:
  shows "A A Le C D"
  using Le_def between_trivial2 cong_trivial_identity by blast

lemma le_cases:
  shows "A B Le C D  C D Le A B"
  by (metis (full_types) cong_reflexivity l5_5_2 l5_6 not_bet_distincts segment_construction_2)

lemma le_zero:
  assumes "A B Le C C"
  shows "A = B"
  by (metis assms cong_diff_4 le_anti_symmetry le_trivial)

lemma le_diff:
  assumes "A  B" and "A B Le C D"
  shows "C  D"
  using assms(1) assms(2) le_zero by blast

lemma lt_diff:
  assumes "A B Lt C D"
  shows "C  D"
  using Lt_def assms cong_trivial_identity le_zero by blast

lemma bet_cong_eq:
  assumes "Bet A B C" and
    "Bet A C D" and
    "Cong B C A D"
  shows "C = D  A = B"
proof -
  have "Bet C B A"
    using Bet_perm assms(1) by blast
  then show ?thesis
    by (metis (no_types) Cong_perm Le_def assms(2) assms(3) between_cong cong_pseudo_reflexivity le_anti_symmetry)
qed

lemma cong__le:
  assumes "Cong A B C D"
  shows "A B Le C D"
  using Le_def assms between_trivial by blast

lemma cong__le3412:
  assumes "Cong A B C D"
  shows "C D Le A B"
  using assms cong__le cong_symmetry by blast

lemma le1221:
  shows "A B Le B A"
  by (simp add: cong__le cong_pseudo_reflexivity)

lemma le_left_comm:
  assumes "A B Le C D"
  shows "B A Le C D"
  using assms le1221 le_transitivity by blast

lemma le_right_comm:
  assumes "A B Le C D"
  shows "A B Le D C"
  by (meson assms cong_right_commutativity l5_5_1 l5_5_2)

lemma le_comm:
  assumes "A B Le C D"
  shows "B A Le D C"
  using assms le_left_comm le_right_comm by blast

lemma ge_left_comm:
  assumes "A B Ge C D"
  shows "B A Ge C D"
  by (meson Ge_def assms le_right_comm)

lemma ge_right_comm:
  assumes "A B Ge C D"
  shows "A B Ge D C"
  using Ge_def assms le_left_comm by presburger

lemma ge_comm0:
  assumes "A B Ge C D"
  shows "B A Ge D C"
  by (meson assms ge_left_comm ge_right_comm)

lemma lt_right_comm:
  assumes "A B Lt C D"
  shows "A B Lt D C"
  using Lt_def assms le_right_comm not_cong_1243 by blast

lemma lt_left_comm:
  assumes "A B Lt C D"
  shows "B A Lt C D"
  using Lt_def assms le_comm lt_right_comm not_cong_2143 by blast

lemma lt_comm:
  assumes "A B Lt C D"
  shows "B A Lt D C"
  using assms lt_left_comm lt_right_comm by blast

lemma gt_left_comm0:
  assumes "A B Gt C D"
  shows "B A Gt C D"
  by (meson Gt_def assms lt_right_comm)

lemma gt_right_comm:
  assumes "A B Gt C D"
  shows "A B Gt D C"
  using Gt_def assms lt_left_comm by presburger

lemma gt_comm:
  assumes "A B Gt C D"
  shows "B A Gt D C"
  by (meson assms gt_left_comm0 gt_right_comm)

lemma cong2_lt__lt:
  assumes "A B Lt C D" and
    "Cong A B A' B'" and
    "Cong C D C' D'"
  shows "A' B' Lt C' D'"
  by (meson Lt_def assms(1) assms(2) assms(3) l5_6 le_anti_symmetry not_cong_3412)

lemma fourth_point:
  assumes "A  B" and
    "B  C" and
    "Col A B P" and
    "Bet A B C"
  shows "Bet P A B  Bet A P B  Bet B P C  Bet B C P"
  by (metis Col_def Tarski_neutral_dimensionless.l5_2 Tarski_neutral_dimensionless_axioms assms(3) assms(4) between_symmetry)

lemma third_point:
  assumes "Col A B P"
  shows "Bet P A B  Bet A P B  Bet A B P"
  using Col_def assms between_symmetry by blast

lemma l5_12_a:
  assumes "Bet A B C"
  shows "A B Le A C  B C Le A C"
  using assms between_symmetry cong_left_commutativity cong_reflexivity l5_5_2 le_left_comm by blast

lemma bet__le1213:
  assumes "Bet A B C"
  shows "A B Le A C"
  using assms l5_12_a by blast

lemma bet__le2313:
  assumes "Bet A B C"
  shows "B C Le A C"
  by (simp add: assms l5_12_a)

lemma bet__lt1213:
  assumes "B  C" and
    "Bet A B C"
  shows "A B Lt A C"
  using Lt_def assms(1) assms(2) bet__le1213 between_cong by blast

lemma bet__lt2313:
  assumes "A  B" and
    "Bet A B C"
  shows "B C Lt A C"
  using Lt_def assms(1) assms(2) bet__le2313 bet_cong_eq l5_1 by blast

lemma l5_12_b:
  assumes "Col A B C" and
    "A B Le A C" and
    "B C Le A C"
  shows "Bet A B C"
  by (metis assms(1) assms(2) assms(3) between_cong col_permutation_5 l5_12_a le_anti_symmetry not_cong_2143 third_point)

lemma bet_le_eq:
  assumes "Bet A B C"
    and "A C Le B C"
  shows "A = B"
  by (meson assms(1) assms(2) bet__le2313 bet_cong_eq l5_1 le_anti_symmetry)

lemma or_lt_cong_gt:
  "A B Lt C D  A B Gt C D  Cong A B C D"
  by (meson Gt_def Lt_def cong_symmetry local.le_cases)

lemma lt__le:
  assumes "A B Lt C D"
  shows "A B Le C D"
  using Lt_def assms by blast

lemma le1234_lt__lt:
  assumes "A B Le C D" and
    "C D Lt E F"
  shows "A B Lt E F"
  by (meson Lt_def assms(1) assms(2) cong__le3412 le_anti_symmetry le_transitivity)

lemma le3456_lt__lt:
  assumes "A B Lt C D" and
    "C D Le E F"
  shows "A B Lt E F"
  by (meson Lt_def assms(1) assms(2) cong2_lt__lt cong_reflexivity le1234_lt__lt)

lemma lt_transitivity:
  assumes "A B Lt C D" and
    "C D Lt E F"
  shows "A B Lt E F"
  using Lt_def assms(1) assms(2) le1234_lt__lt by blast

lemma not_and_lt:
  "¬ (A B Lt C D  C D Lt A B)"
  by (simp add: Lt_def le_anti_symmetry)

lemma nlt:
  "¬ A B Lt A B"
  using not_and_lt by blast

lemma le__nlt:
  assumes "A B Le C D"
  shows "¬ C D Lt A B"
  using assms le3456_lt__lt nlt by blast

lemma cong__nlt:
  assumes "Cong A B C D"
  shows "¬ A B Lt C D"
  by (simp add: Lt_def assms)

lemma nlt__le:
  assumes "¬ A B Lt C D"
  shows "C D Le A B"
  using Lt_def assms cong__le3412 local.le_cases by blast

lemma lt__nle:
  assumes "A B Lt C D"
  shows "¬ C D Le A B"
  using assms le__nlt by blast

lemma nle__lt:
  assumes "¬ A B Le C D"
  shows "C D Lt A B"
  using assms nlt__le by blast

lemma lt1123:
  assumes "B  C"
  shows "A A Lt B C"
  using assms le_diff nle__lt by blast

lemma bet2_le2__le_R1:
  assumes "Bet a P b" and
    "Bet A Q B" and
    "P a Le Q A" and
    "P b Le Q B" and
    "B = Q"
  shows "a b Le A B"
  by (metis assms(3) assms(4) assms(5) le_comm le_diff)

lemma bet2_le2__le_R2:
  assumes "Bet a Po b" and
    "Bet A PO B" and
    "Po a Le PO A" and
    "Po b Le PO B" and
    "A  PO" and
    "B  PO"
  shows "a b Le A B"
proof -
  obtain b' where P1: "Bet A PO b'  Cong PO b' b Po"
    using segment_construction by blast
  obtain a' where P2: "Bet B PO a'  Cong PO a' a Po"
    using segment_construction by blast
  obtain a'' where P3: "Bet PO a'' A  Cong Po a PO a''"
    using Le_def assms(3) by blast
  have P4: "a' = a''"
    by (meson Bet_cases P2 P3 assms(2) assms(6) between_inner_transitivity cong_right_commutativity construction_uniqueness not_cong_3412)
  have P5: "B a' Le B A"
    using Bet_cases P3 P4 assms(2) bet__le1213 between_exchange2 by blast
  obtain b'' where P6: "Bet PO b'' B  Cong Po b PO b''"
    using Le_def assms(4) by blast
  then have "b' = b''"
    using P1 assms(2) assms(5) between_inner_transitivity cong_right_commutativity construction_uniqueness not_cong_3412 by blast
  then have "a' b' Le a' B"
    using Bet_cases P2 P6 bet__le1213 between_exchange2 by blast
  then have "a' b' Le A B"
    using P5 le_comm le_transitivity by blast
  thus ?thesis
    by (smt Cong_cases P1 P3 P4 Tarski_neutral_dimensionless.l5_6 Tarski_neutral_dimensionless_axioms assms(1) between_exchange3 between_symmetry cong_reflexivity l2_11_b)
qed

lemma bet2_le2__le:
  assumes "Bet a P b" and
    "Bet A Q B" and
    "P a Le Q A" and
    "P b Le Q B"
  shows "a b Le A B"
proof cases
  assume "A = Q"
  thus ?thesis
    using assms(3) assms(4) le_diff by force
next
  assume "¬ A = Q"
  thus ?thesis
    using assms(1) assms(2) assms(3) assms(4) bet2_le2__le_R1 bet2_le2__le_R2 by blast
qed

lemma Le_cases:
  assumes "A B Le C D  B A Le C D  A B Le D C  B A Le D C"
  shows "A B Le C D"
  using assms le_left_comm le_right_comm by blast

lemma Lt_cases:
  assumes "A B Lt C D  B A Lt C D  A B Lt D C  B A Lt D C"
  shows "A B Lt C D"
  using assms lt_comm lt_left_comm by blast

subsection "Out lines"

lemma bet_out:
  assumes "B  A" and
    "Bet A B C"
  shows  "A Out B C"
  using Out_def assms(1) assms(2) bet_neq12__neq by fastforce

lemma bet_out_1:
  assumes "B  A" and
    "Bet C B A"
  shows "A Out B C"
  by (simp add: assms(1) assms(2) bet_out between_symmetry)

lemma out_dec:
  shows "P Out A B  ¬ P Out A B"
  by simp

lemma out_diff1:
  assumes "A Out B C"
  shows "B  A"
  using Out_def assms by auto

lemma out_diff2:
  assumes "A Out B C"
  shows "C  A"
  using Out_def assms by auto

lemma out_distinct:
  assumes "A Out B C"
  shows  "B  A  C  A"
  using assms out_diff1 out_diff2 by auto

lemma out_col:
  assumes "A Out B C"
  shows "Col A B C"
  using Col_def Out_def assms between_symmetry by auto

lemma l6_2:
  assumes "A  P" and
    "B  P" and
    "C  P" and
    "Bet A P C"
  shows "Bet B P C  P Out A B"
  by (smt Bet_cases Out_def assms(1) assms(2) assms(3) assms(4) between_inner_transitivity l5_2 outer_transitivity_between)

lemma bet_out__bet:
  assumes "Bet A P C" and
    "P Out A B"
  shows "Bet B P C"
  by (metis Tarski_neutral_dimensionless.l6_2 Tarski_neutral_dimensionless_axioms assms(1) assms(2) not_bet_distincts out_diff1)

lemma l6_3_1:
  assumes "P Out A B"
  shows "A  P  B  P  ( C. (C  P  Bet A P C  Bet B P C))"
  using assms bet_out__bet out_diff1 out_diff2 point_construction_different by fastforce

lemma l6_3_2:
  assumes "A  P" and
    "B  P" and
    " C. (C  P  Bet A P C  Bet B P C)"
  shows "P Out A B"
  using assms(1) assms(2) assms(3) l6_2 by blast

lemma l6_4_1:
  assumes "P Out A B" and
    "Col A P B"
  shows "¬ Bet A P B"
  using Out_def assms(1) between_equality between_symmetry by fastforce

lemma l6_4_2:
  assumes "Col A P B"
    and "¬ Bet A P B"
  shows "P Out A B"
  by (metis Out_def assms(1) assms(2) bet_out col_permutation_1 third_point)

lemma out_trivial:
  assumes "A  P"
  shows "P Out A A"
  by (simp add: assms bet_out_1 between_trivial2)

lemma l6_6:
  assumes "P Out A B"
  shows "P Out B A"
  using Out_def assms by auto

lemma l6_7:
  assumes "P Out A B" and
    "P Out B C"
  shows "P Out A C"
  by (smt Out_def assms(1) assms(2) between_exchange4 l5_1 l5_3)

lemma bet_out_out_bet:
  assumes "Bet A B C" and
    "B Out A A'" and
    "B Out C C'"
  shows "Bet A' B C'"
  by (metis Out_def assms(1) assms(2) assms(3) bet_out__bet between_inner_transitivity outer_transitivity_between)

lemma out2_bet_out:
  assumes "B Out A C" and
    "B Out X P" and
    "Bet A X C"
  shows "B Out A P  B Out C P"
  by (smt Out_def Tarski_neutral_dimensionless.l6_7 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) between_exchange2 between_symmetry)

lemma l6_11_uniqueness:
  assumes "A Out X R" and
    "Cong A X B C" and
    "A Out Y R" and
    "Cong A Y B C"
  shows "X = Y"
  by (metis Out_def assms(1) assms(2) assms(3) assms(4) between_cong cong_symmetry cong_transitivity l6_6 l6_7)

lemma l6_11_existence:
  assumes "R  A" and
    "B  C"
  shows " X. (A Out X R  Cong A X B C)"
  by (metis Out_def assms(1) assms(2) cong_reverse_identity segment_construction_2)


lemma segment_construction_3:
  assumes "A  B" and
    "X  Y"
  shows " C. (A Out B C  Cong A C X Y)"
  by (metis assms(1) assms(2) l6_11_existence l6_6)

lemma l6_13_1:
  assumes "P Out A B" and
    "P A Le P B"
  shows "Bet P A B"
  by (metis Out_def assms(1) assms(2) bet__lt1213 le__nlt)

lemma l6_13_2:
  assumes "P Out A B" and
    "Bet P A B"
  shows "P A Le P B"
  by (simp add: assms(2) bet__le1213)

lemma l6_16_1:
  assumes "P  Q" and
    "Col S P Q" and
    "Col X P Q"
  shows "Col X P S"
  by (smt Col_def assms(1) assms(2) assms(3) bet3__bet col_permutation_4 l5_1 l5_3 outer_transitivity_between third_point)

lemma col_transitivity_1:
  assumes "P  Q" and
    "Col P Q A" and
    "Col P Q B"
  shows "Col P A B"
  by (meson Tarski_neutral_dimensionless.l6_16_1 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) not_col_permutation_2)

lemma col_transitivity_2:
  assumes "P  Q" and
    "Col P Q A" and
    "Col P Q B"
  shows "Col Q A B"
  by (metis Tarski_neutral_dimensionless.col_transitivity_1 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) not_col_permutation_4)

lemma l6_21:
  assumes "¬ Col A B C" and
    "C  D" and
    "Col A B P" and
    "Col A B Q" and
    "Col C D P" and
    "Col C D Q"
  shows "P = Q"
  by (metis assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) col_transitivity_1 l6_16_1 not_col_distincts)

lemma col2__eq:
  assumes "Col A X Y" and
    "Col B X Y" and
    "¬ Col A X B"
  shows  "X = Y"
  using assms(1) assms(2) assms(3) l6_16_1 by blast

lemma not_col_exists:
  assumes  "A  B"
  shows " C. ¬ Col A B C"
  by (metis Col_def assms col_transitivity_2 lower_dim_ex)

lemma col3:
  assumes "X  Y" and
    "Col X Y A" and
    "Col X Y B" and
    "Col X Y C"
  shows "Col A B C"
  by (metis assms(1) assms(2) assms(3) assms(4) col_transitivity_2)

lemma colx:
  assumes "A  B" and
    "Col X Y A" and
    "Col X Y B" and
    "Col A B C"
  shows "Col X Y C"
  by (metis assms(1) assms(2) assms(3) assms(4) l6_21 not_col_distincts)

lemma out2__bet:
  assumes "A Out B C" and
    "C Out A B"
  shows "Bet A B C"
  by (metis Out_def assms(1) assms(2) between_equality between_symmetry)

lemma bet2_le2__le1346:
  assumes "Bet A B C" and
    "Bet A' B' C'" and
    "A B Le A' B'" and
    "B C Le B' C'"
  shows "A C Le A' C'"
  using Le_cases assms(1) assms(2) assms(3) assms(4) bet2_le2__le by blast

lemma bet2_le2__le2356_R1:
  assumes "Bet A A C" and
    "Bet A' B' C'" and
    "A A Le A' B'" and
    "A' C' Le A C"
  shows "B' C' Le A C"
  using assms(2) assms(4) bet__le2313 le3456_lt__lt lt__nle nlt__le by blast

lemma bet2_le2__le2356_R2:
  assumes "A  B" and
    "Bet A B C" and
    "Bet A' B' C'" and
    "A B Le A' B'" and
    "A' C' Le A C"
  shows "B' C' Le B C"
proof -
  have "A  C"
    using assms(1) assms(2) bet_neq12__neq by blast
  obtain B0 where P1: "Bet A B B0  Cong A B0 A' B'"
    using assms(4) l5_5_1 by blast
  then have P2: "A  B0"
    using assms(1) bet_neq12__neq by blast
  obtain C0 where P3: "Bet A C0 C  Cong A' C' A C0"
    using Le_def assms(5) by blast
  then have "A  C0"
    using assms(1) assms(3) assms(4) bet_neq12__neq cong_diff le_diff by blast
  then have P4: "Bet A B0 C0"
    by (smt Out_def P1 P2 P3 assms(1) assms(2) assms(3) bet__le1213 between_exchange2 between_symmetry l5_1 l5_3 l5_6 l6_13_1 not_cong_3412)
  have K1: "B0 C0 Le B C0"
    using P1 P4 between_exchange3 l5_12_a by blast
  have K2: "B C0 Le B C"
    using P1 P3 P4 bet__le1213 between_exchange3 between_exchange4 by blast
  then have "Cong B0 C0 B' C'"
    using P1 P3 P4 assms(3) l4_3_1 not_cong_3412 by blast
  then show ?thesis
    by (meson K1 K2 cong__nlt le_transitivity nlt__le)
qed

lemma bet2_le2__le2356:
  assumes "Bet A B C" and
    "Bet A' B' C'" and
    "A B Le A' B'" and
    "A' C' Le A C"
  shows "B' C' Le B C"
proof (cases)
  assume "A = B"
  then show ?thesis
    using assms(1) assms(2) assms(3) assms(4) bet2_le2__le2356_R1 by blast
next
  assume "¬ A = B"
  then show ?thesis
    using assms(1) assms(2) assms(3) assms(4) bet2_le2__le2356_R2 by blast
qed

lemma bet2_le2__le1245:
  assumes "Bet A B C" and
    "Bet A' B' C'" and
    "B C Le B' C'" and
    "A' C' Le A C"
  shows "A' B' Le A B"
  using assms(1) assms(2) assms(3) assms(4) bet2_le2__le2356 between_symmetry le_comm by blast

lemma cong_preserves_bet:
  assumes "Bet B A' A0" and
    "Cong B A' E D'" and
    "Cong B A0 E D0" and
    "E Out D' D0"
  shows "Bet E D' D0"
  using Tarski_neutral_dimensionless.l6_13_1 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) assms(4) bet__le1213 l5_6 by fastforce

lemma out_cong_cong:
  assumes "B Out A A0" and
    "E Out D D0" and
    "Cong B A E D" and
    "Cong B A0 E D0"
  shows "Cong A A0 D D0"
  by (meson Out_def assms(1) assms(2) assms(3) assms(4) cong_4321 cong_symmetry l4_3_1 l5_6 l6_13_1 l6_13_2)

lemma not_out_bet:
  assumes "Col A B C" and
    "¬ B Out A C"
  shows "Bet A B C"
  using assms(1) assms(2) l6_4_2 by blast

lemma or_bet_out:
  shows "Bet A B C  B Out A C  ¬ Col A B C"
  using not_out_bet by blast

lemma not_bet_out:
  assumes "Col A B C" and
    "¬ Bet A B C"
  shows "B Out A C"
  by (simp add: assms(1) assms(2) l6_4_2)

lemma not_bet_and_out:
  shows  "¬ (Bet A B C  B Out A C)"
  using bet_col l6_4_1 by blast

lemma out_to_bet:
  assumes "Col A' B' C'" and
    "B Out A C  B' Out A' C'" and
    "Bet A B C"
  shows "Bet A' B' C'"
  using assms(1) assms(2) assms(3) not_bet_and_out or_bet_out by blast

lemma col_out2_col:
  assumes "Col A B C" and
    "B Out A AA" and
    "B Out C CC"
  shows "Col AA B CC" using l6_21
  by (smt Tarski_neutral_dimensionless.out_col Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) col_trivial_2 not_col_permutation_1 out_diff1)

lemma bet2_out_out:
  assumes "B  A" and
    "B'  A" and
    "A Out C C'" and
    "Bet A B C" and
    "Bet A B' C'"
  shows "A Out B B'"
  by (meson assms(1) assms(2) assms(3) assms(4) assms(5) bet_out l6_6 l6_7)

lemma bet2__out:
  assumes "A  B" and
    "A  B'" and
    "Bet A B C"
    and "Bet A B' C"
  shows "A Out B B'"
  using Out_def assms(1) assms(2) assms(3) assms(4) l5_3 by auto

lemma out_bet_out_1:
  assumes "P Out A C" and
    "Bet A B C"
  shows "P Out A B"
  by (metis assms(1) assms(2) not_bet_and_out out2_bet_out out_trivial)

lemma out_bet_out_2:
  assumes "P Out A C" and
    "Bet A B C"
  shows "P Out B C"
  using assms(1) assms(2) l6_6 l6_7 out_bet_out_1 by blast

lemma out_bet__out:
  assumes "Bet P Q A" and
    "Q Out A B"
  shows "P Out A B"
  by (smt Out_def assms(1) assms(2) bet_out_1 bet_out__bet)

lemma segment_reverse:
  assumes "Bet A B C "
  shows " B'. Bet A B' C  Cong C B' A B"
  by (metis Bet_perm Cong_perm assms bet_cong_eq cong_reflexivity segment_construction_2)

lemma diff_col_ex:
  shows " C. A  C  B  C  Col A B C"
  by (metis bet_col bet_neq12__neq point_construction_different)

lemma diff_bet_ex3:
  assumes "Bet A B C"
  shows " D. A  D  B  D  C  D  Col A B D"
  by (metis (mono_tags, opaque_lifting) Col_def bet_out_1 between_trivial2 col_transitivity_1 l6_4_1 point_construction_different)

lemma diff_col_ex3:
  assumes "Col A B C"
  shows " D. A  D  B  D  C  D  Col A B D"
  by (metis Bet_perm Col_def between_equality between_trivial2 point_construction_different)

lemma Out_cases:
  assumes "A Out B C  A Out C B"
  shows "A Out B C"
  using assms l6_6 by blast

subsection "Midpoint"

lemma midpoint_dec:
  "I Midpoint A B  ¬ I Midpoint A B"
  by simp

lemma is_midpoint_id:
  assumes "A Midpoint A B"
  shows "A = B"
  using Midpoint_def assms between_cong by blast

lemma is_midpoint_id_2:
  assumes "A Midpoint B A"
  shows "A = B"
  using Midpoint_def assms cong_diff_2 by blast

lemma l7_2:
  assumes "M Midpoint A B"
  shows "M Midpoint B A"
  using Bet_perm Cong_perm Midpoint_def assms by blast

lemma l7_3:
  assumes "M Midpoint A A"
  shows "M = A"
  using Midpoint_def assms bet_neq23__neq by blast

lemma l7_3_2:
  "A Midpoint A A"
  by (simp add: Midpoint_def between_trivial2 cong_reflexivity)

lemma symmetric_point_construction:
  " P'. A Midpoint P P'"
  by (meson Midpoint_def cong__le cong__le3412 le_anti_symmetry segment_construction)

lemma symmetric_point_uniqueness:
  assumes "P Midpoint A P1" and
    "P Midpoint A P2"
  shows "P1 = P2"
  by (metis Midpoint_def assms(1) assms(2) between_cong_3 cong_diff_4 cong_inner_transitivity)

lemma l7_9:
  assumes "A Midpoint P X" and
    "A Midpoint Q X"
  shows "P = Q"
  using assms(1) assms(2) l7_2 symmetric_point_uniqueness by blast

lemma l7_9_bis:
  assumes "A Midpoint P X" and
    "A Midpoint X Q"
  shows "P = Q"
  using assms(1) assms(2) l7_2 symmetric_point_uniqueness by blast

lemma l7_13_R1:
  assumes "A  P" and
    "A Midpoint P' P" and
    "A Midpoint Q' Q"
  shows "Cong P Q P' Q'"
proof -
  obtain X where P1: "Bet P' P X  Cong P X Q A"
    using segment_construction by blast
  obtain X' where P2: "Bet X P' X'  Cong P' X' Q A"
    using segment_construction by blast
  obtain Y where P3: "Bet Q' Q Y  Cong Q Y P A"
    using segment_construction by blast
  obtain Y' where P4: "Bet Y Q' Y'  Cong Q' Y' P A"
    using segment_construction by blast
  have P5: "Bet Y A Q'"
    by (meson Midpoint_def P3 P4 assms(3) bet3__bet between_symmetry l5_3)
  have P6: "Bet P' A X"
    using Midpoint_def P1 assms(2) between_exchange4 by blast
  have P7: "Bet A P X"
    using Midpoint_def P1 assms(2) between_exchange3 by blast
  have P8: "Bet Y Q A"
    using Midpoint_def P3 assms(3) between_exchange3 between_symmetry by blast
  have P9: "Bet A Q' Y'"
    using P4 P5 between_exchange3 by blast
  have P10: "Bet X' P' A"
    using P2 P6 between_exchange3 between_symmetry by blast
  have P11: "Bet X A X'"
    using P10 P2 P6 between_symmetry outer_transitivity_between2 by blast
  have P12: "Bet Y A Y'"
    using P4 P5 between_exchange4 by blast
  have P13: "Cong A X Y A"
    using P1 P3 P7 P8 l2_11_b not_cong_4321 by blast
  have P14: "Cong A Y' X' A"
  proof -
    have Q1: "Cong Q' Y' P' A"
      using Midpoint_def P4 assms(2) cong_transitivity not_cong_3421 by blast
    have "Cong A Q' X' P'"
      by (meson Midpoint_def P2 assms(3) cong_transitivity not_cong_3421)
    then show ?thesis
      using P10 P9 Q1 l2_11_b by blast
  qed
  have P15: "Cong A Y A Y'"
  proof -
    have "Cong Q Y Q' Y'"
      using P3 P4 cong_transitivity not_cong_3412 by blast
    then show ?thesis
      using Bet_perm Cong_perm Midpoint_def P8 P9 assms(3) l2_11_b by blast
  qed
  have P16: "Cong X A Y' A"
    using Cong_cases P13 P15 cong_transitivity by blast
  have P17: "Cong A X' A Y"
    using P14 P15 cong_transitivity not_cong_3421 by blast
  have P18: "X A X' Y' FSC Y' A Y X"
  proof -
    have Q3: "Col X A X'"
      by (simp add: Col_def P11)
    have "Cong X X' Y' Y"
      using Bet_cases P11 P12 P16 P17 l2_11_b by blast
    then show ?thesis
      by (simp add: Cong3_def FSC_def P16 P17 Q3 cong_4321 cong_pseudo_reflexivity)
  qed
  then have "Y Q A X IFSC Y' Q' A X'"
    by (smt IFSC_def Midpoint_def P14 P15 P16 P7 P8 P9 assms(1) assms(3) bet_neq12__neq between_symmetry cong_4321 cong_inner_transitivity cong_right_commutativity l4_16)
  then have "X P A Q IFSC X' P' A Q'"
    by (meson IFSC_def Midpoint_def P10 P7 assms(2) between_symmetry cong_4312 l4_2)
  then show ?thesis
    using l4_2 by auto
qed

lemma l7_13:
  assumes "A Midpoint P' P" and
    "A Midpoint Q' Q"
  shows "Cong P Q P' Q'"
proof (cases)
  assume "A = P"
  then show ?thesis
    using Midpoint_def assms(1) assms(2) cong_3421 is_midpoint_id_2 by blast
next
  show ?thesis
    by (metis Tarski_neutral_dimensionless.l7_13_R1 Tarski_neutral_dimensionless_axioms assms(1) assms(2) cong_trivial_identity is_midpoint_id_2 not_cong_2143)
qed

lemma l7_15:
  assumes "A Midpoint P P'" and
    "A Midpoint Q Q'" and
    "A Midpoint R R'" and
    "Bet P Q R"
  shows "Bet P' Q' R'"
proof -
  have "P Q R Cong3 P' Q' R'"
    using Cong3_def assms(1) assms(2) assms(3) l7_13 l7_2 by blast
  then show ?thesis
    using assms(4) l4_6 by blast
qed

lemma l7_16:
  assumes "A Midpoint P P'" and
    "A Midpoint Q Q'" and
    "A Midpoint R R'" and
    "A Midpoint S S'" and
    "Cong P Q R S"
  shows "Cong P' Q' R' S'"
  by (meson assms(1) assms(2) assms(3) assms(4) assms(5) cong_transitivity l7_13 not_cong_3412)

lemma symmetry_preserves_midpoint:
  assumes "Z Midpoint A D" and
    "Z Midpoint B E" and
    "Z Midpoint C F" and
    "B Midpoint A C"
  shows "E Midpoint D F"
  by (meson Midpoint_def assms(1) assms(2) assms(3) assms(4) l7_15 l7_16)

lemma Mid_cases:
  assumes "A Midpoint B C  A Midpoint C B"
  shows "A Midpoint B C"
  using assms l7_2 by blast

lemma Mid_perm:
  assumes "A Midpoint B C"
  shows "A Midpoint B C  A Midpoint C B"
  by (simp add: assms l7_2)

lemma l7_17:
  assumes "A Midpoint P P'" and
    "B Midpoint P P'"
  shows "A = B"
proof -
  obtain pp :: "'p  'p  'p" where
    f1: "p pa. p Midpoint pa (pp p pa)"
    by (meson symmetric_point_construction)
  then have "p pa. Bet pa p (pp p pa)"
    by (meson Midpoint_def)
  then have f2: "p. Bet p p p"
    by (meson between_inner_transitivity)
  have f3: "p pa. Bet (pp pa p) pa p"
    using f1 Mid_perm Midpoint_def by blast
  have f4: "p. pp p p = p"
    using f2 f1 by (metis Midpoint_def bet_cong_eq)
  have f5: "Bet (pp P P') P B"
    using f3 by (meson Midpoint_def assms(2) between_inner_transitivity)
  have f6: "A Midpoint P' P"
    using Mid_perm assms(1) by blast
  have f7: "Bet (pp P P') P A"
    using f3 Midpoint_def assms(1) between_inner_transitivity by blast
  have f8: "Bet P' A P"
    using f6 by (simp add: Midpoint_def)
  have "Cong P' A A P"
    using f6 Midpoint_def by blast
  then have "P' = P  A = B"
    using f8 by (metis (no_types) Midpoint_def assms(2) bet_cong_eq between_inner_transitivity l5_2)
  then show ?thesis
    using f7 f6 f5 f4 f1 by (metis (no_types) Col_perm Mid_perm assms(2) bet_col l4_18 l5_2 l7_13)
qed

lemma l7_17_bis:
  assumes "A Midpoint P P'" and
    "B Midpoint P' P"
  shows "A = B"
  by (meson Tarski_neutral_dimensionless.l7_17 Tarski_neutral_dimensionless.l7_2 Tarski_neutral_dimensionless_axioms assms(1) assms(2))

lemma l7_20:
  assumes "Col A M B" and
    "Cong M A M B"
  shows "A = B  M Midpoint A B"
  by (metis Bet_cases Col_def Midpoint_def assms(1) assms(2) between_cong cong_left_commutativity not_cong_3412)

lemma l7_20_bis:
  assumes "A  B" and
    "Col A M B" and
    "Cong M A M B"
  shows "M Midpoint A B"
  using assms(1) assms(2) assms(3) l7_20 by blast

lemma cong_col_mid:
  assumes "A  C" and
    "Col A B C" and
    "Cong A B B C"
  shows "B Midpoint A C"
  using assms(1) assms(2) assms(3) cong_left_commutativity l7_20 by blast

lemma l7_21_R1:
  assumes "¬ Col A B C" and
    "B  D" and
    "Cong A B C D" and
    "Cong B C D A" and
    "Col A P C" and
    "Col B P D"
  shows "P Midpoint A C"
proof -
  obtain X where P1: "B D P Cong3 D B X"
    using Col_perm assms(6) cong_pseudo_reflexivity l4_14 by blast
  have P2: "Col D B X"
    using P1 assms(6) l4_13 not_col_permutation_5 by blast
  have P3: "B D P A FSC D B X C"
    using FSC_def P1 assms(3) assms(4) assms(6) not_col_permutation_5 not_cong_2143 not_cong_3412 by blast
  have P4: "B D P C FSC D B X A"
    by (simp add: FSC_def P1 assms(3) assms(4) assms(6) col_permutation_5 cong_4321)
  have "A P C Cong3 C X A"
    using Cong3_def Cong_perm P3 P4 assms(2) cong_pseudo_reflexivity l4_16 by blast
  then show ?thesis
    by (smt Cong3_def NCol_cases P2 assms(1) assms(2) assms(5) assms(6) colx cong_col_mid l4_13 not_col_distincts not_col_permutation_1 not_cong_1243)
qed

lemma l7_21:
  assumes "¬ Col A B C" and
    "B  D" and
    "Cong A B C D" and
    "Cong B C D A" and
    "Col A P C" and
    "Col B P D"
  shows "P Midpoint A C  P Midpoint B D"
  by (smt assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) col_transitivity_2 is_midpoint_id_2 l7_21_R1 not_col_distincts not_cong_3412)

lemma l7_22_aux_R1:
  assumes "Bet A1 C C" and
    "Bet B1 C B2" and
    "Cong C A1 C B1" and
    "Cong C C C B2" and
    "M1 Midpoint A1 B1" and
    "M2 Midpoint A2 B2"and
    "C A1 Le C C"
  shows "Bet M1 C M2"
  by (metis assms(3) assms(5) assms(7) cong_diff_3 l7_3 le_diff not_bet_distincts)

lemma l7_22_aux_R2:
  assumes "A2  C" and
    "Bet A1 C A2" and
    "Bet B1 C B2" and
    "Cong C A1 C B1" and
    "Cong C A2 C B2" and
    "M1 Midpoint A1 B1" and
    "M2 Midpoint A2 B2" and
    "C A1 Le C A2"
  shows "Bet M1 C M2"
proof -
  obtain X where P1: "C Midpoint A2 X"
    using symmetric_point_construction by blast
  obtain X0 where P2: "C Midpoint B2 X0"
    using symmetric_point_construction by blast
  obtain X1 where P3: "C Midpoint M2 X1"
    using symmetric_point_construction by blast
  have P4: "X1 Midpoint X X0"
    using P1 P2 P3 assms(7) symmetry_preserves_midpoint by blast
  have P5: "C A1 Le C X"
    using Cong_perm Midpoint_def P1 assms(8) cong_reflexivity l5_6 by blast
  have P6: "Bet C A1 X"
    by (smt Midpoint_def P1 P5 assms(1) assms(2) bet2__out between_symmetry is_midpoint_id_2 l5_2 l6_13_1)
  have P7: "C B1 Le C X0"
  proof -
    have Q1: "Cong C A1 C B1"
      by (simp add: assms(4))
    have "Cong C X C X0"
      using P1 P2 assms(5) l7_16 l7_3_2 by blast
    then show ?thesis
      using P5 Q1 l5_6 by blast
  qed
  have P8: "Bet C B1 X0"
    by (smt Midpoint_def P2 P7 assms(1) assms(3) assms(5) bet2__out between_symmetry cong_identity l5_2 l6_13_1)
  obtain Q where P9: "Bet X1 Q C  Bet A1 Q B1"
    by (meson Bet_perm Midpoint_def P4 P6 P8 l3_17)
  have P10: "X A1 C X1 IFSC X0 B1 C X1"
    by (smt Cong_perm IFSC_def Midpoint_def P1 P2 P4 P6 P8 assms(4) assms(5) between_symmetry cong_reflexivity l7_16 l7_3_2)
  have P11: "Cong A1 X1 B1 X1"
    using P10 l4_2 by blast
  have P12: "Cong Q A1 Q B1"
  proof (cases)
    assume "C = X1"
    then show ?thesis
      using P9 assms(4) bet_neq12__neq by blast
  next
    assume Q1: "¬ C = X1"
    have Q2: "Col C X1 Q"
      using Col_def P9 by blast
    have Q3: "Cong C A1 C B1"
      by (simp add: assms(4))
    have "Cong X1 A1 X1 B1"
      using P11 not_cong_2143 by blast
    then show ?thesis
      using Q1 Q2 Q3 l4_17 by blast
  qed
  have P13: "Q Midpoint A1 B1"
    by (simp add: Midpoint_def P12 P9 cong_left_commutativity)
  then show ?thesis
    using Midpoint_def P3 P9 assms(6) between_inner_transitivity between_symmetry l7_17 by blast
qed

lemma l7_22_aux:
  assumes "Bet A1 C A2" and
    "Bet B1 C B2" and
    "Cong C A1 C B1" and
    "Cong C A2 C B2" and
    "M1 Midpoint A1 B1" and
    "M2 Midpoint A2 B2" and
    "C A1 Le C A2"
  shows "Bet M1 C M2"
  by (smt assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) l7_22_aux_R1 l7_22_aux_R2)

lemma l7_22:
  assumes "Bet A1 C A2" and
    "Bet B1 C B2" and
    "Cong C A1 C B1" and
    "Cong C A2 C B2" and
    "M1 Midpoint A1 B1" and
    "M2 Midpoint A2 B2"
  shows "Bet M1 C M2"
  by (meson assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) between_symmetry l7_22_aux local.le_cases)

lemma bet_col1:
  assumes "Bet A B D" and
    "Bet A C D"
  shows "Col A B C"
  using Bet_perm Col_def assms(1) assms(2) l5_3 by blast

lemma l7_25_R1:
  assumes "Cong C A C B" and
    "Col A B C"
  shows " X. X Midpoint A B"
  using assms(1) assms(2) l7_20 l7_3_2 not_col_permutation_5 by blast

lemma l7_25_R2:
  assumes "Cong C A C B" and
    "¬ Col A B C"
  shows " X. X Midpoint A B"
proof -
  obtain P where P1: "Bet C A P  A  P"
    using point_construction_different by auto
  obtain Q where P2: "Bet C B Q  Cong B Q A P"
    using segment_construction by blast
  obtain R where P3: "Bet A R Q  Bet B R P"
    using P1 P2 between_symmetry inner_pasch by blast
  obtain X where P4: "Bet A X B  Bet R X C"
    using P1 P3 inner_pasch by blast
  have "Cong X A X B"
  proof -
    have Q1: "Cong R A R B  Cong X A X B"
    proof (cases)
      assume "R = C"
      then show ?thesis
        using P4 bet_neq12__neq by blast
    next
      assume Q2: "¬ R = C"
      have "Col R C X"
        using Col_perm P4 bet_col by blast
      then show ?thesis
        using Q2 assms(1) l4_17 by blast
    qed
    have "Cong R A R B"
    proof -
      have Q3: "C A P B OFSC C B Q A"
        by (simp add: OFSC_def P1 P2 assms(1) cong_pseudo_reflexivity cong_symmetry)
      have Q4: "Cong P B Q A"
        using Q3 assms(2) five_segment_with_def not_col_distincts by blast
      obtain R' where Q5: "Bet A R' Q  B R P Cong3 A R' Q"
        using Cong_perm P3 Q4 l4_5 by blast
      have Q6: "B R P A IFSC A R' Q B"
        by (meson Cong3_def IFSC_def OFSC_def P3 Q3 Q5 not_cong_2143)
      have Q7: "B R P Q IFSC A R' Q P"
        using IFSC_def P2 Q6 cong_pseudo_reflexivity by auto
      have Q8: "Cong R A R' B"
        using Q6 l4_2 by auto
      have Q9: "Cong R Q R' P"
        using Q7 l4_2 by auto
      have Q10: "A R Q Cong3 B R' P"
        using Cong3_def Q4 Q8 Q9 cong_commutativity not_cong_4321 by blast
      have Q11: "Col B R' P"
        using P3 Q10 bet_col l4_13 by blast
      have "R = R'"
      proof -
        have R1: "B  P"
          using P1 assms(1) between_cong by blast
        then have R2: "A  Q"
          using Q4 cong_diff_2 by blast
        have R3: "B  Q"
          using P1 P2 cong_diff_3 by blast
        then have R4: "B  R"
          by (metis Cong3_def P1 Q11 Q5 assms(2) bet_col cong_diff_3 l6_21 not_col_distincts)
        have R5: "¬ Col A Q B"
          by (metis P2 R3 assms(2) bet_col col_permutation_3 col_trivial_2 l6_21)
        have R6: "B  P"
          by (simp add: R1)
        have R7: "Col A Q R"
          using NCol_cases P3 bet_col by blast
        have R8: "Col A Q R'"
          using Q5 bet_col col_permutation_5 by blast
        have R9: "Col B P R"
          using NCol_cases P3 bet_col by blast
        have  "Col B P R'"
          using Col_perm Q11 by blast
        then show ?thesis
          using R5 R6 R7 R8 R9 l6_21 by blast
      qed
      then show ?thesis
        using Q8 by blast
    qed
    then show ?thesis
      using Q1 by blast
  qed
  then show ?thesis
    using P4 assms(2) bet_col l7_20_bis not_col_distincts by blast
qed

lemma l7_25:
  assumes "Cong C A C B"
  shows " X. X Midpoint A B"
  using assms l7_25_R1 l7_25_R2 by blast

lemma midpoint_distinct_1:
  assumes "A  B" and
    "I Midpoint A B"
  shows "I  A  I  B"
  using assms(1) assms(2) is_midpoint_id is_midpoint_id_2 by blast

lemma midpoint_distinct_2:
  assumes "I  A" and
    "I Midpoint A B"
  shows "A  B  I  B"
  using assms(1) assms(2) is_midpoint_id_2 l7_3 by blast

lemma midpoint_distinct_3:
  assumes "I  B" and
    "I Midpoint A B"
  shows "A  B  I  A"
  using assms(1) assms(2) is_midpoint_id l7_3 by blast

lemma midpoint_def:
  assumes "Bet A B C" and
    "Cong A B B C"
  shows "B Midpoint A C"
  using Midpoint_def assms(1) assms(2) by blast

lemma midpoint_bet:
  assumes "B Midpoint A C"
  shows "Bet A B C"
  using Midpoint_def assms by blast

lemma midpoint_col:
  assumes "M Midpoint A B"
  shows "Col M A B"
  using assms bet_col col_permutation_4 midpoint_bet by blast

lemma midpoint_cong:
  assumes "B Midpoint A C"
  shows "Cong A B B C"
  using Midpoint_def assms by blast

lemma midpoint_out:
  assumes "A  C" and
    "B Midpoint A C"
  shows "A Out B C"
  using assms(1) assms(2) bet_out midpoint_bet midpoint_distinct_1 by blast

lemma midpoint_out_1:
  assumes "A  C" and
    "B Midpoint A C"
  shows "C Out A B"
  by (metis Tarski_neutral_dimensionless.midpoint_bet Tarski_neutral_dimensionless.midpoint_distinct_1 Tarski_neutral_dimensionless_axioms assms(1) assms(2) bet_out_1 l6_6)

lemma midpoint_not_midpoint:
  assumes "A  B" and
    "I Midpoint A B"
  shows "¬ B Midpoint A I"
  using assms(1) assms(2) between_equality_2 midpoint_bet midpoint_distinct_1 by blast

lemma swap_diff:
  assumes "A  B"
  shows "B  A"
  using assms by auto

lemma cong_cong_half_1:
  assumes "M Midpoint A B" and
    "M' Midpoint A' B'" and
    "Cong A B A' B'"
  shows "Cong A M A' M'"
proof -
  obtain M'' where P1: "Bet A' M'' B'  A M B Cong3 A' M'' B'"
    using assms(1) assms(3) l4_5 midpoint_bet by blast
  have P2: "M'' Midpoint A' B'"
    by (meson Cong3_def P1 assms(1) cong_inner_transitivity midpoint_cong midpoint_def)
  have P3: "M' = M''"
    using P2 assms(2) l7_17 by auto
  then show ?thesis
    using Cong3_def P1 by blast
qed

lemma cong_cong_half_2:
  assumes "M Midpoint A B" and
    "M' Midpoint A' B'" and
    "Cong A B A' B'"
  shows "Cong B M B' M'"
  using assms(1) assms(2) assms(3) cong_cong_half_1 l7_2 not_cong_2143 by blast

lemma cong_mid2__cong:
  assumes "M Midpoint A B" and
    "M' Midpoint A' B'" and
    "Cong A M A' M'"
  shows "Cong A B A' B'"
  by (meson assms(1) assms(2) assms(3) cong_inner_transitivity l2_11_b midpoint_bet midpoint_cong)

lemma mid__lt:
  assumes "A  B" and
    "M Midpoint A B"
  shows "A M Lt A B"
  using assms(1) assms(2) bet__lt1213 midpoint_bet midpoint_distinct_1 by blast

lemma le_mid2__le13:
  assumes "M Midpoint A B" and
    "M' Midpoint A' B'" and
    "A M Le A' M'"
  shows "A B Le A' B'"
  by (smt Tarski_neutral_dimensionless.cong_mid2__cong Tarski_neutral_dimensionless.l7_13 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) bet2_le2__le2356 l5_6 l7_3_2 le_anti_symmetry le_comm local.le_cases midpoint_bet)

lemma le_mid2__le12:
  assumes "M Midpoint A B" and
    "M' Midpoint A' B'"
    and "A B Le A' B'"
  shows "A M Le A' M'"
  by (meson assms(1) assms(2) assms(3) cong__le3412 cong_cong_half_1 le_anti_symmetry le_mid2__le13 local.le_cases)

lemma lt_mid2__lt13:
  assumes "M Midpoint A B" and
    "M' Midpoint A' B'" and
    "A M Lt A' M'"
  shows "A B Lt A' B'"
  by (meson Tarski_neutral_dimensionless.le_mid2__le12 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) lt__nle nlt__le)

lemma lt_mid2__lt12:
  assumes "M Midpoint A B" and
    "M' Midpoint A' B'" and
    "A B Lt A' B'"
  shows "A M Lt A' M'"
  by (meson Tarski_neutral_dimensionless.le_mid2__le13 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) le__nlt nle__lt)

lemma midpoint_preserves_out:
  assumes "A Out B C" and
    "M Midpoint A A'" and
    "M Midpoint B B'" and
    "M Midpoint C C'"
  shows "A' Out B' C'"
  by (smt Out_def assms(1) assms(2) assms(3) assms(4) l6_4_2 l7_15 l7_2 not_bet_and_out not_col_distincts)

lemma col_cong_bet:
  assumes "Col A B D" and
    "Cong A B C D" and
    "Bet A C B"
  shows "Bet C A D  Bet C B D"
  by (smt Col_def assms(1) assms(2) assms(3) bet_cong_eq between_inner_transitivity col_transitivity_2 cong_4321 l6_2 not_bet_and_out not_cong_4312 third_point)

lemma col_cong2_bet1:
  assumes "Col A B D" and
    "Bet A C B" and
    "Cong A B C D" and
    "Cong A C B D"
  shows "Bet C B D"
  by (metis assms(1) assms(2) assms(3) assms(4) bet__le1213 bet_cong_eq between_symmetry col_cong_bet cong__le cong_left_commutativity l5_12_b l5_6 outer_transitivity_between2)

lemma col_cong2_bet2:
  assumes "Col A B D" and
    "Bet A C B" and
    "Cong A B C D" and
    "Cong A D B C"
  shows "Bet C A D"
  by (metis assms(1) assms(2) assms(3) assms(4) bet_cong_eq col_cong_bet cong_identity not_bet_distincts not_cong_3421 outer_transitivity_between2)

lemma col_cong2_bet3:
  assumes "Col A B D" and
    "Bet A B C" and
    "Cong A B C D" and
    "Cong A C B D"
  shows "Bet B C D"
  by (metis assms(1) assms(2) assms(3) assms(4) bet__le1213 bet__le2313 bet_col col_transitivity_2 cong_diff_3 cong_reflexivity l5_12_b l5_6 not_bet_distincts)

lemma col_cong2_bet4:
  assumes "Col A B C" and
    "Bet A B D" and
    "Cong A B C D" and
    "Cong A D B C"
  shows "Bet B D C"
  using assms(1) assms(2) assms(3) assms(4) col_cong2_bet3 cong_right_commutativity by blast

lemma col_bet2_cong1:
  assumes "Col A B D" and
    "Bet A C B" and
    "Cong A B C D" and
    "Bet C B D"
  shows "Cong A C D B"
  by (meson assms(2) assms(3) assms(4) between_symmetry cong_pseudo_reflexivity cong_right_commutativity l4_3)

lemma col_bet2_cong2:
  assumes "Col A B D" and
    "Bet A C B" and
    "Cong A B C D" and
    "Bet C A D"
  shows "Cong D A B C"
  by (meson assms(2) assms(3) assms(4) between_symmetry cong_commutativity cong_pseudo_reflexivity cong_symmetry l4_3)

lemma bet2_lt2__lt:
  assumes "Bet a Po b" and
    "Bet A PO B" and
    "Po a Lt PO A" and
    "Po b Lt PO B"
  shows "a b Lt A B"
  by (metis Lt_cases Tarski_neutral_dimensionless.nle__lt Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) assms(4) bet2_le2__le1245 le__nlt lt__le)

lemma bet2_lt_le__lt:
  assumes "Bet a Po b" and
    "Bet A PO B" and
    "Cong Po a PO A" and
    "Po b Lt PO B"
  shows "a b Lt A B"
  by (smt Lt_def Tarski_neutral_dimensionless.l4_3_1 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) assms(4) bet2_le2__le cong__le not_cong_2143)

subsection "Orthogonality"

lemma per_dec:
  "Per A B C  ¬ Per A B C"
  by simp

lemma l8_2:
  assumes "Per A B C"
  shows "Per C B A"
proof -
  obtain C' where P1: "B Midpoint C C'  Cong A C A C'"
    using Per_def assms by blast
  obtain A' where P2: "B Midpoint A A'"
    using symmetric_point_construction by blast
  have "Cong C' A C A'"
    using Mid_perm P1 P2 l7_13 by blast
  thus ?thesis
    using P1 P2 Per_def cong_4321 cong_inner_transitivity by blast
qed

lemma Per_cases:
  assumes "Per A B C  Per C B A"
  shows "Per A B C"
  using assms l8_2 by blast

lemma Per_perm :
  assumes "Per A B C"
  shows "Per A B C  Per C B A"
  by (simp add: assms l8_2)

lemma l8_3 :
  assumes "Per A B C" and
    "A  B" and
    "Col B A A'"
  shows "Per A' B C"
  by (smt Per_def assms(1) assms(2) assms(3) l4_17 l7_13 l7_2 l7_3_2)

lemma l8_4:
  assumes "Per A B C" and
    "B Midpoint C C'"
  shows "Per A B C'"
  by (metis Tarski_neutral_dimensionless.l8_2 Tarski_neutral_dimensionless_axioms assms(1) assms(2) l8_3 midpoint_col midpoint_distinct_1)

lemma l8_5:
  "Per A B B"
  using Per_def cong_reflexivity l7_3_2 by blast

lemma l8_6:
  assumes "Per A B C" and
    "Per A' B C" and
    "Bet A C A'"
  shows "B = C"
  by (metis Per_def assms(1) assms(2) assms(3) l4_19 midpoint_distinct_3 symmetric_point_uniqueness)

lemma l8_7:
  assumes "Per A B C" and
    "Per A C B"
  shows "B = C"
proof -
  obtain C' where P1: "B Midpoint C C'  Cong A C A C'"
    using Per_def assms(1) by blast
  obtain A' where P2: "C Midpoint A A'"
    using Per_def assms(2) l8_2 by blast
  have "Per C' C A"
    by (metis P1 Tarski_neutral_dimensionless.l8_3 Tarski_neutral_dimensionless_axioms assms(2) bet_col l8_2 midpoint_bet midpoint_distinct_3)
  then have "Cong A C' A' C'"
    using Cong_perm P2 Per_def symmetric_point_uniqueness by blast
  then have "Cong A' C A' C'"
    using P1 P2 cong_inner_transitivity midpoint_cong not_cong_2134 by blast
  then have Q4: "Per A' B C"
    using P1 Per_def by blast
  have "Bet A' C A"
    using Mid_perm P2 midpoint_bet by blast
  thus ?thesis
    using Q4 assms(1) l8_6 by blast
qed

lemma l8_8:
  assumes "Per A B A"
  shows "A = B"
  using Tarski_neutral_dimensionless.l8_6 Tarski_neutral_dimensionless_axioms assms between_trivial2 by fastforce

lemma per_distinct:
  assumes "Per A B C" and
    "A  B"
  shows "A  C"
  using assms(1) assms(2) l8_8 by blast

lemma per_distinct_1:
  assumes "Per A B C" and
    "B  C"
  shows "A  C"
  using assms(1) assms(2) l8_8 by blast

lemma l8_9:
  assumes "Per A B C" and
    "Col A B C"
  shows "A = B  C = B"
  using Col_cases assms(1) assms(2) l8_3 l8_8 by blast

lemma l8_10:
  assumes "Per A B C" and
    "A B C Cong3 A' B' C'"
  shows "Per A' B' C'"
proof -
  obtain D where P1: "B Midpoint C D  Cong A C A D"
    using Per_def assms(1) by blast
  obtain D' where P2: "Bet C' B' D'  Cong B' D' B' C'"
    using segment_construction by blast
  have P3: "B' Midpoint C' D'"
    by (simp add: Midpoint_def P2 cong_4312)
  have "Cong A' C' A' D'"
  proof (cases)
    assume "C = B"
    thus ?thesis
      by (metis Cong3_def P3 assms(2) cong_diff_4 cong_reflexivity is_midpoint_id)
  next
    assume Q1: "¬ C = B"
    have "C B D A OFSC C' B' D' A'"
      by (metis Cong3_def OFSC_def P1 P3 Tarski_neutral_dimensionless.cong_mid2__cong Tarski_neutral_dimensionless_axioms assms(2) cong_commutativity l4_3_1 midpoint_bet)
    thus ?thesis
      by (meson OFSC_def P1 Q1 cong_4321 cong_inner_transitivity five_segment_with_def)
  qed
  thus ?thesis
    using Per_def P3 by blast
qed

lemma col_col_per_per:
  assumes "A  X" and
    "C  X" and
    "Col U A X" and
    "Col V C X" and
    "Per A X C"
  shows "Per U X V"
  by (meson Tarski_neutral_dimensionless.l8_2 Tarski_neutral_dimensionless.l8_3 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) assms(4) assms(5) not_col_permutation_3)

lemma perp_in_dec:
  "X PerpAt A B C D  ¬ X PerpAt A B C D"
  by simp

lemma perp_distinct:
  assumes "A B Perp C D"
  shows "A  B  C  D"
  using PerpAt_def Perp_def assms by auto

lemma l8_12:
  assumes "X PerpAt A B C D"
  shows "X PerpAt C D A B"
  using Per_perm PerpAt_def assms by auto

lemma per_col:
  assumes "B  C" and
    "Per A B C" and
    "Col B C D"
  shows "Per A B D"
  by (metis Tarski_neutral_dimensionless.l8_3 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) l8_2)

lemma l8_13_2:
  assumes "A  B" and
    "C  D" and
    "Col X A B" and
    "Col X C D" and
    " U.  V. Col U A B  Col V C D  U  X  V  X  Per U X V"
  shows "X PerpAt A B C D"
proof -
  obtain pp :: 'p and ppa :: 'p where
    f1: "Col pp A B  Col ppa C D  pp  X  ppa  X  Per pp X ppa"
    using assms(5) by blast
  obtain ppb :: "'p  'p  'p  'p  'p  'p" and ppc :: "'p  'p  'p  'p  'p  'p" where
    "x0 x1 x2 x3 x4. (v5 v6. (Col v5 x3 x2  Col v6 x1 x0)  ¬ Per v5 x4 v6) = ((Col (ppb x0 x1 x2 x3 x4) x3 x2  Col (ppc x0 x1 x2 x3 x4) x1 x0)  ¬ Per (ppb x0 x1 x2 x3 x4) x4 (ppc x0 x1 x2 x3 x4))"
    by moura
  then have f2: "p pa pb pc pd. (¬ p PerpAt pa pb pc pd  pa  pb  pc  pd  Col p pa pb  Col p pc pd  (pe pf. (¬ Col pe pa pb  ¬ Col pf pc pd)  Per pe p pf))  (p PerpAt pa pb pc pd  pa = pb  pc = pd  ¬ Col p pa pb  ¬ Col p pc pd  (Col (ppb pd pc pb pa p) pa pb  Col (ppc pd pc pb pa p) pc pd)  ¬ Per (ppb pd pc pb pa p) p (ppc pd pc pb pa p))"
    using PerpAt_def by fastforce
  { assume "¬ Col (ppb D C B A X) pp X"
    then have "¬ Col (ppb D C B A X) A B  ¬ Col (ppc D C B A X) C D  Per (ppb D C B A X) X (ppc D C B A X)"
      using f1 by (meson assms(1) assms(3) col3 not_col_permutation_2) }
  moreover
  { assume "¬ Col (ppc D C B A X) ppa X"
    then have "¬ Col (ppb D C B A X) A B  ¬ Col (ppc D C B A X) C D  Per (ppb D C B A X) X (ppc D C B A X)"
      using f1 by (meson assms(2) assms(4) col3 not_col_permutation_2) }
  ultimately have "¬ Col (ppb D C B A X) A B  ¬ Col (ppc D C B A X) C D  Per (ppb D C B A X) X (ppc D C B A X)"
    using f1 by (meson Tarski_neutral_dimensionless.col_col_per_per Tarski_neutral_dimensionless_axioms)
  then have "(X PerpAt A B C D  A = B  C = D  ¬ Col X A B  ¬ Col X C D  Col (ppb D C B A X) A B  Col (ppc D C B A X) C D  ¬ Per (ppb D C B A X) X (ppc D C B A X))  (¬ Col (ppb D C B A X) A B  ¬ Col (ppc D C B A X) C D  Per (ppb D C B A X) X (ppc D C B A X))"
    using f2 by presburger
  thus ?thesis
    using assms(1) assms(2) assms(3) assms(4) by blast
qed

lemma l8_14_1:
  "¬ A B Perp A B"
  by (metis PerpAt_def Perp_def Tarski_neutral_dimensionless.col_trivial_1 Tarski_neutral_dimensionless.col_trivial_3 Tarski_neutral_dimensionless_axioms l8_8)

lemma l8_14_2_1a:
  assumes "X PerpAt A B C D"
  shows "A B Perp C D"
  using Perp_def assms by blast

lemma perp_in_distinct:
  assumes "X PerpAt A B C D"
  shows "A  B  C  D"
  using PerpAt_def assms by blast

lemma l8_14_2_1b:
  assumes "X PerpAt A B C D" and
    "Col Y A B" and
    "Col Y C D"
  shows "X = Y"
  by (metis PerpAt_def assms(1) assms(2) assms(3) l8_13_2 l8_14_1 l8_14_2_1a)

lemma l8_14_2_1b_bis:
  assumes "A B Perp C D" and
    "Col X A B" and
    "Col X C D"
  shows "X PerpAt A B C D"
  using Perp_def assms(1) assms(2) assms(3) l8_14_2_1b by blast

lemma l8_14_2_2:
  assumes "A B Perp C D" and
    " Y. (Col Y A B  Col Y C D)  X = Y"
  shows "X PerpAt A B C D"
  by (metis Tarski_neutral_dimensionless.PerpAt_def Tarski_neutral_dimensionless.Perp_def Tarski_neutral_dimensionless_axioms assms(1) assms(2))

lemma l8_14_3:
  assumes "X PerpAt A B C D" and
    "Y PerpAt A B C D"
  shows "X = Y"
  by (meson PerpAt_def assms(1) assms(2) l8_14_2_1b)

lemma l8_15_1:
  assumes "Col A B X" and
    "A B Perp C X"
  shows "X PerpAt A B C X"
  using NCol_perm assms(1) assms(2) col_trivial_3 l8_14_2_1b_bis by blast

lemma l8_15_2:
  assumes "Col A B X" and
    "X PerpAt A B C X"
  shows "A B Perp C X"
  using assms(2) l8_14_2_1a by blast

lemma perp_in_per:
  assumes "B PerpAt A B B C"
  shows "Per A B C"
  by (meson NCol_cases PerpAt_def assms col_trivial_3)

lemma perp_sym:
  assumes "A B Perp A B"
  shows "C D Perp C D"
  using assms l8_14_1 by auto

lemma perp_col0:
  assumes "A B Perp C D" and
    "X  Y" and
    "Col A B X" and
    "Col A B Y"
  shows "C D Perp X Y"
proof -
  obtain X0 where P1: "X0 PerpAt A B C D"
    using Perp_def assms(1) by blast
  then have P2: " A  B  C  D  Col X0 A B  Col X0 C D 
  ((Col U A B  Col V C D)  Per U X0 V)" using PerpAt_def by blast
  have Q1: "C  D" using P2 by blast
  have Q2: "X  Y" using assms(2) by blast
  have Q3: "Col X0 C D" using P2 by blast
  have Q4: "Col X0 X Y"
  proof -
    have "p pa. Col p pa Y  Col p pa X  Col p pa X0  p  pa"
      by (metis (no_types) Col_cases P2 assms(3) assms(4))
    thus ?thesis
      using col3 by blast
  qed
  have "X0 PerpAt C D X Y"
  proof -
    have " U V. (Col U C D  Col V X Y)  Per U X0 V"
      by (metis Col_perm P1 Per_perm Q2 Tarski_neutral_dimensionless.PerpAt_def Tarski_neutral_dimensionless_axioms assms(3) assms(4) colx)
    thus ?thesis using Q1 Q2 Q3 Q4 PerpAt_def by blast
  qed
  thus ?thesis
    using Perp_def by auto
qed

lemma per_perp_in:
  assumes "A  B" and
    "B  C" and
    "Per A B C"
  shows "B PerpAt A B B C"
  by (metis Col_def assms(1) assms(2) assms(3) between_trivial2 l8_13_2)

lemma per_perp:
  assumes "A  B" and
    "B  C" and
    "Per A B C"
  shows "A B Perp B C"
  using Perp_def assms(1) assms(2) assms(3) per_perp_in by blast

lemma perp_left_comm:
  assumes "A B Perp C D"
  shows "B A Perp C D"
proof -
  obtain X where "X PerpAt A B C D"
    using Perp_def assms by blast
  then have "X PerpAt B A C D"
    using PerpAt_def col_permutation_5 by auto
  thus ?thesis
    using Perp_def by blast
qed

lemma perp_right_comm:
  assumes "A B Perp C D"
  shows "A B Perp D C"
  by (meson Perp_def assms l8_12 perp_left_comm)

lemma perp_comm:
  assumes "A B Perp C D"
  shows "B A Perp D C"
  by (simp add: assms perp_left_comm perp_right_comm)

lemma perp_in_sym:
  assumes  "X PerpAt A B C D"
  shows "X PerpAt C D A B"
  by (simp add: assms l8_12)

lemma perp_in_left_comm:
  assumes "X PerpAt A B C D"
  shows "X PerpAt B A C D"
  by (metis Col_cases PerpAt_def assms)

lemma perp_in_right_comm:
  assumes "X PerpAt A B C D"
  shows "X PerpAt A B D C"
  using assms perp_in_left_comm perp_in_sym by blast

lemma perp_in_comm:
  assumes "X PerpAt A B C D"
  shows "X PerpAt B A D C"
  by (simp add: assms perp_in_left_comm perp_in_right_comm)

lemma Perp_cases:
  assumes "A B Perp C D  B A Perp C D  A B Perp D C  B A Perp D C  C D Perp A B  C D Perp B A  D C Perp A B  D C Perp B A"
  shows "A B Perp C D"
  by (meson Perp_def assms perp_in_sym perp_left_comm)

lemma Perp_perm :
  assumes "A B Perp C D"
  shows "A B Perp C D  B A Perp C D  A B Perp D C  B A Perp D C  C D Perp A B  C D Perp B A  D C Perp A B  D C Perp B A"
  by (meson Perp_def assms perp_in_sym perp_left_comm)

lemma Perp_in_cases:
  assumes "X PerpAt A B C D  X PerpAt B A C D  X PerpAt A B D C  X PerpAt B A D C  X PerpAt C D A B  X PerpAt C D B A  X PerpAt D C A B  X PerpAt D C B A"
  shows "X PerpAt A B C D"
  using assms perp_in_left_comm perp_in_sym by blast

lemma Perp_in_perm:
  assumes "X PerpAt A B C D"
  shows "X PerpAt A B C D  X PerpAt B A C D  X PerpAt A B D C  X PerpAt B A D C  X PerpAt C D A B  X PerpAt C D B A  X PerpAt D C A B  X PerpAt D C B A"
  using Perp_in_cases assms by blast

lemma perp_in_col:
  assumes "X PerpAt A B C D"
  shows "Col A B X  Col C D X"
  using PerpAt_def assms col_permutation_2 by presburger

lemma perp_perp_in:
  assumes "A B Perp C A"
  shows "A PerpAt A B C A"
  using assms l8_15_1 not_col_distincts by blast

lemma perp_per_1:
  assumes "A B Perp C A"
  shows "Per B A C"
  using Perp_in_cases assms perp_in_per perp_perp_in by blast

lemma perp_per_2:
  assumes "A B Perp A C"
  shows "Per B A C"
  by (simp add: Perp_perm assms perp_per_1)

lemma perp_col:
  assumes "A  E" and
    "A B Perp C D" and
    "Col A B E"
  shows "A E Perp C D"
  using Perp_perm assms(1) assms(2) assms(3) col_trivial_3 perp_col0 by blast

lemma perp_col2:
  assumes "A B Perp X Y" and
    "C  D" and
    "Col A B C" and
    "Col A B D"
  shows "C D Perp X Y"
  using Perp_perm assms(1) assms(2) assms(3) assms(4) perp_col0 by blast

lemma perp_col4:
  assumes "P  Q" and
    "R  S" and
    "Col A B P" and
    "Col A B Q" and
    "Col C D R" and
    "Col C D S" and
    "A B Perp C D"
  shows "P Q Perp R S"
  using assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) perp_col0 by blast

lemma perp_not_eq_1:
  assumes "A B Perp C D"
  shows "A  B"
  using assms perp_distinct by auto

lemma perp_not_eq_2:
  assumes "A B Perp C D"
  shows "C  D"
  using assms perp_distinct by auto

lemma diff_per_diff:
  assumes "A  B" and
    "Cong A P B R" and
    "Per B A P"
    and "Per A B R"
  shows "P  R"
  using assms(1) assms(3) assms(4) l8_2 l8_7 by blast

lemma per_not_colp:
  assumes "A  B" and
    "A  P" and
    "B  R" and
    "Per B A P"
    and "Per A B R"
  shows "¬ Col P A R"
  by (metis Per_cases Tarski_neutral_dimensionless.col_permutation_4 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(4) assms(5) l8_3 l8_7)

lemma per_not_col:
  assumes "A  B" and
    "B  C" and
    "Per A B C"
  shows "¬ Col A B C"
  using assms(1) assms(2) assms(3) l8_9 by auto

lemma perp_not_col2:
  assumes "A B Perp C D"
  shows "¬ Col A B C  ¬ Col A B D"
  using assms l8_14_1 perp_col2 perp_distinct by blast

lemma perp_not_col:
  assumes "A B Perp P A"
  shows "¬ Col A B P"
proof -
  have "A PerpAt A B P A"
    using assms perp_perp_in by auto
  then have "Per P A B"
    by (simp add: perp_in_per perp_in_sym)
  then have "¬ Col B A P"
    by (metis NCol_perm Tarski_neutral_dimensionless.perp_not_eq_1 Tarski_neutral_dimensionless.perp_not_eq_2 Tarski_neutral_dimensionless_axioms assms per_not_col)
  thus ?thesis
    using Col_perm by blast
qed

lemma perp_in_col_perp_in:
  assumes "C  E" and
    "Col C D E" and
    "P PerpAt A B C D"
  shows  "P PerpAt A B C E"
proof -
  have P2: "C  D"
    using assms(3) perp_in_distinct by blast
  have P3: "Col P A B"
    using PerpAt_def assms(3) by auto
  have "Col P C D"
    using PerpAt_def assms(3) by blast
  then have "Col P C E"
    using P2 assms(2) col_trivial_2 colx by blast
  thus ?thesis
    by (smt P3 Perp_perm Tarski_neutral_dimensionless.l8_14_2_1b_bis Tarski_neutral_dimensionless.perp_col Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) l8_14_2_1a)
qed

lemma perp_col2_bis:
  assumes "A B Perp C D" and
    "Col C D P" and
    "Col C D Q" and
    "P  Q"
  shows "A B Perp P Q"
  using Perp_cases assms(1) assms(2) assms(3) assms(4) perp_col0 by blast

lemma perp_in_perp_bis_R1:
  assumes "X  A" and
    "X PerpAt A B C D"
  shows "X B Perp C D  A X Perp C D"
  by (metis assms(2) l8_14_2_1a perp_col perp_in_col)

lemma perp_in_perp_bis:
  assumes "X PerpAt A B C D"
  shows "X B Perp C D  A X Perp C D"
  by (metis assms l8_14_2_1a perp_in_perp_bis_R1)

lemma col_per_perp:
  assumes "A  B" and
    "B  C" and
    (*   "D ≠ B" and  *)
    "D  C" and
    "Col B C D" and
    "Per A B C"
  shows "C D Perp A B"
  by (metis Perp_cases assms(1) assms(2) assms(3) assms(4) assms(5) col_trivial_2 per_perp perp_col2_bis)

lemma per_cong_mid_R1:
  assumes "B = H" and
    (*  "B ≠ C" and *)
    "Bet A B C" and
    "Cong A H C H" and
    "Per H B C"
  shows "B Midpoint A C"
  using assms(1) assms(2) assms(3) midpoint_def not_cong_1243 by blast

lemma per_cong_mid_R2:
  assumes (*"B ≠ H" and *)
    "B  C" and
    "Bet A B C" and
    "Cong A H C H" and
    "Per H B C"
  shows "B Midpoint A C"
proof -
  have P1: "Per C B H"
    using Per_cases assms(4) by blast
  have P2: "Per H B A"
    using assms(1) assms(2) assms(4) bet_col col_permutation_1 per_col by blast
  then have P3: "Per A B H"
    using Per_cases by blast
  obtain C' where P4: "B Midpoint C C'  Cong H C H C'"
    using Per_def assms(4) by blast
  obtain H' where P5: "B Midpoint H H'  Cong C H C H'"
    using P1 Per_def by blast
  obtain A' where P6: "B Midpoint A A'  Cong H A H A'"
    using P2 Per_def by blast
  obtain H'' where P7: "B Midpoint H H''  Cong A H A H'"
    using P3 P5 Tarski_neutral_dimensionless.Per_def Tarski_neutral_dimensionless_axioms symmetric_point_uniqueness by fastforce
  then have P8: "H' = H''"
    using P5 symmetric_point_uniqueness by blast
  have "H B H' A IFSC H B H' C"
  proof -
    have Q1: "Bet H B H'"
      by (simp add: P7 P8 midpoint_bet)
    have Q2: "Cong H H' H H'"
      by (simp add: cong_reflexivity)
    have Q3: "Cong B H' B H'"
      by (simp add: cong_reflexivity)
    have Q4: "Cong H A H C"
      using assms(3) not_cong_2143 by blast
    have "Cong H' A H' C"
      using P5 P7 assms(3) cong_commutativity cong_inner_transitivity by blast
    thus ?thesis
      by (simp add: IFSC_def Q1 Q2 Q3 Q4)
  qed
  thus ?thesis
    using assms(1) assms(2) bet_col bet_neq23__neq l4_2 l7_20_bis by auto
qed

lemma per_cong_mid:
  assumes "B  C" and
    "Bet A B C" and
    "Cong A H C H" and
    "Per H B C"
  shows "B Midpoint A C"
  using assms(1) assms(2) assms(3) assms(4) per_cong_mid_R1 per_cong_mid_R2 by blast

lemma per_double_cong:
  assumes "Per A B C" and
    "B Midpoint C C'"
  shows "Cong A C A C'"
  using Mid_cases Per_def assms(1) assms(2) l7_9_bis by blast

lemma cong_perp_or_mid_R1:
  assumes "Col A B X" and
    "A  B" and
    "M Midpoint A B" and
    "Cong A X B X"
  shows "X = M  ¬ Col A B X  M PerpAt X M A B"
  using assms(1) assms(2) assms(3) assms(4) col_permutation_5 cong_commutativity l7_17_bis l7_2 l7_20 by blast

lemma cong_perp_or_mid_R2:
  assumes "¬ Col A B X" and
    "A  B" and
    "M Midpoint A B" and
    "Cong A X B X"
  shows "X = M  ¬ Col A B X  M PerpAt X M A B"
proof -
  have P1: "Col M A B"
    by (simp add: assms(3) midpoint_col)
  have "Per X M A"
    using Per_def assms(3) assms(4) cong_commutativity by blast
  thus ?thesis
    by (metis P1 assms(1) assms(2) assms(3) midpoint_distinct_1 not_col_permutation_4 per_perp_in perp_in_col_perp_in perp_in_right_comm)
qed

lemma cong_perp_or_mid:
  assumes "A  B" and
    "M Midpoint A B" and
    "Cong A X B X"
  shows "X = M  ¬ Col A B X  M PerpAt X M A B"
  using assms(1) assms(2) assms(3) cong_perp_or_mid_R1 cong_perp_or_mid_R2 by blast

lemma col_per2_cases:
  assumes "B  C" and
    "B'  C" and
    "C  D" and
    "Col B C D" and
    "Per A B C" and
    "Per A B' C"
  shows  "B = B'  ¬ Col B' C D"
  by (meson Tarski_neutral_dimensionless.l8_7 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) l6_16_1 per_col)

lemma l8_16_1:
  assumes "Col A B X" and
    "Col A B U" and
    "A B Perp C X"
  shows "¬ Col A B C  Per C X U"
  by (metis assms(1) assms(2) assms(3) l8_5 perp_col0 perp_left_comm perp_not_col2 perp_per_2)

lemma l8_16_2:
  assumes "Col A B X" and
    "Col A B U"
    and "U  X" and
    "¬ Col A B C" and
    "Per C X U"
  shows "A B Perp C X"
proof -
  obtain X where "X PerpAt A B C X"
    by (metis (no_types) NCol_perm assms(1) assms(2) assms(3) assms(4) assms(5) l8_13_2 l8_2 not_col_distincts)
  thus ?thesis
    by (smt Perp_perm assms(1) assms(2) assms(3) assms(4) assms(5) col3 col_per_perp not_col_distincts per_col per_perp)
qed

lemma l8_18_uniqueness:
  assumes (*"¬ Col A B C" and *)
    "Col A B X" and
    "A B Perp C X" and
    "Col A B Y" and
    "A B Perp C Y"
  shows "X = Y"
  using assms(1) assms(2) assms(3) assms(4) l8_16_1 l8_7 by blast

lemma midpoint_distinct:
  assumes "¬ Col A B C" and
    "Col A B X" and
    "X Midpoint C C'"
  shows "C  C'"
  using assms(1) assms(2) assms(3) l7_3 by auto

lemma l8_20_1_R1:
  assumes "A = B"
  shows "Per B A P"
  by (simp add: assms l8_2 l8_5)

lemma l8_20_1_R2:
  assumes "A  B" and
    "Per A B C" and
    "P Midpoint C' D" and
    "A Midpoint C' C" and
    "B Midpoint D C"
  shows "Per B A P"
proof -
  obtain B' where P1: "A Midpoint B B'"
    using symmetric_point_construction by blast
  obtain D' where P2: "A Midpoint D D'"
    using symmetric_point_construction by blast
  obtain P' where P3: "A Midpoint P P'"
    using symmetric_point_construction by blast
  have P4: "Per B' B C"
    by (metis P1 Tarski_neutral_dimensionless.Per_cases Tarski_neutral_dimensionless.per_col Tarski_neutral_dimensionless_axioms assms(1) assms(2) midpoint_col not_col_permutation_4)
  have P5: "Per B B' C'"
  proof -
    have "Per B' B C"
      by (simp add: P4)
    have "B' B C Cong3 B B' C'"
      by (meson Cong3_def P1 assms(4) l7_13 l7_2)
    thus ?thesis
      using P4 l8_10 by blast
  qed
  have P6: "B' Midpoint D' C'"
    by (meson P1 P2 assms(4) assms(5) l7_15 l7_16 l7_2 midpoint_bet midpoint_cong midpoint_def)
  have P7: "P' Midpoint C D'"
    using P2 P3 assms(3) assms(4) symmetry_preserves_midpoint by blast
  have P8: "A Midpoint P P'"
    by (simp add: P3)
  obtain D'' where P9: "B Midpoint C D''  Cong B' C B' D"
    using P4 assms(5) l7_2 per_double_cong by blast
  have P10: "D'' = D"
    using P9 assms(5) l7_9_bis by blast
  obtain D'' where P11: "B' Midpoint C' D''  Cong B C' B D''"
    using P5 Per_def by blast
  have P12: "D' = D''"
    by (meson P11 P6 Tarski_neutral_dimensionless.l7_9_bis Tarski_neutral_dimensionless_axioms)
  have P13: "P Midpoint C' D"
    using assms(3) by blast
  have P14: "Cong C D C' D'"
    using P2 assms(4) l7_13 l7_2 by blast
  have P15: "Cong C' D C D'"
    using P2 assms(4) cong_4321 l7_13 by blast
  have P16: "Cong P D P' D'"
    using P2 P8 cong_symmetry l7_13 by blast
  have P17: "Cong P D P' C"
    using P16 P7 cong_3421 cong_transitivity midpoint_cong by blast
  have P18: "C' P D B IFSC D' P' C B"
    by (metis Bet_cases IFSC_def P10 P11 P12 P13 P15 P17 P7 P9 cong_commutativity cong_right_commutativity l7_13 l7_3_2 midpoint_bet)
  then have "Cong B P B P'"
    using Tarski_neutral_dimensionless.l4_2 Tarski_neutral_dimensionless_axioms not_cong_2143 by fastforce
  thus ?thesis
    using P8 Per_def by blast
qed

lemma l8_20_1:
  assumes "Per A B C" and
    "P Midpoint C' D" and
    "A Midpoint C' C" and
    "B Midpoint D C"
  shows "Per B A P"
  using assms(1) assms(2) assms(3) assms(4) l8_20_1_R1 l8_20_1_R2 by fastforce

lemma l8_20_2:
  assumes "P Midpoint C' D" and
    "A Midpoint C' C" and
    "B Midpoint D C" and
    "B  C"
  shows "A  P"
  using assms(1) assms(2) assms(3) assms(4) l7_3 symmetric_point_uniqueness by blast

lemma perp_col1:
  assumes "C  X" and
    "A B Perp C D" and
    "Col C D X"
  shows "A B Perp C X"
  using assms(1) assms(2) assms(3) col_trivial_3 perp_col2_bis by blast

lemma l8_18_existence:
  assumes "¬ Col A B C"
  shows " X. Col A B X  A B Perp C X"
proof -
  obtain Y where P1: "Bet B A Y  Cong A Y A C"
    using segment_construction by blast
  then obtain P where P2: "P Midpoint C Y"
    using Mid_cases l7_25 by blast
  then have P3: "Per A P Y"
    using P1 Per_def l7_2 by blast
  obtain Z where P3: "Bet A Y Z  Cong Y Z Y P"
    using segment_construction by blast
  obtain Q where P4: "Bet P Y Q  Cong Y Q Y A"
    using segment_construction by blast
  obtain Q' where P5: "Bet Q Z Q'  Cong Z Q' Q Z"
    using segment_construction by blast
  then have P6: "Z Midpoint Q Q'"
    using midpoint_def not_cong_3412 by blast
  obtain C' where P7: "Bet Q' Y C'  Cong Y C' Y C"
    using segment_construction by blast
  obtain X where P8: "X Midpoint C C'"
    using Mid_cases P7 l7_25 by blast
  have P9: "A Y Z Q OFSC Q Y P A"
    by (simp add: OFSC_def P3 P4 between_symmetry cong_4321 cong_pseudo_reflexivity)
  have P10: "A  Y"
    using P1 assms cong_reverse_identity not_col_distincts by blast
  then have P11: "Cong Z Q P A"
    using P9 five_segment_with_def by blast
  then have P12: "A P Y Cong3 Q Z Y"
    using Cong3_def P3 P4 not_cong_4321 by blast
  have P13: "Per Q Z Y"
    using Cong_perm P1 P12 P2 Per_def l8_10 l8_4 by blast
  then have P14: "Per Y Z Q"
    by (simp add: l8_2)
  have P15: "P  Y"
    using NCol_cases P1 P2 assms bet_col l7_3_2 l7_9_bis by blast
  obtain Q'' where P16:"Z Midpoint Q Q''  Cong Y Q Y Q'"
    using P14 P6 per_double_cong by blast
  then have P17: "Q' = Q''"
    using P6 symmetric_point_uniqueness by blast
  have P18: "Bet Z Y X"
  proof -
    have "Bet Q Y C"
      using P15 P2 P4 between_symmetry midpoint_bet outer_transitivity_between2 by blast
    thus ?thesis
      using P16 P6 P7 P8 l7_22 not_cong_3412 by blast
  qed
  have P19: "Q  Y"
    using P10 P4 cong_reverse_identity by blast
  have P20: "Per Y X C"
  proof -
    have "Bet C P Y"
      by (simp add: P2 midpoint_bet)
    thus ?thesis
      using P7 P8 Per_def not_cong_3412 by blast
  qed
  have P21: "Col P Y Q"
    by (simp add: Col_def P4)
  have P22: "Col P Y C"
    using P2 midpoint_col not_col_permutation_5 by blast
  have P23: "Col P Q C"
    using P15 P21 P22 col_transitivity_1 by blast
  have P24: "Col Y Q C"
    using P15 P21 P22 col_transitivity_2 by auto
  have P25: "Col A Y B"
    by (simp add: Col_def P1)
  have P26: "Col A Y Z"
    using P3 bet_col by blast
  have P27: "Col A B Z"
    using P10 P25 P26 col_transitivity_1 by blast
  have P28: "Col Y B Z"
    using P10 P25 P26 col_transitivity_2 by blast
  have P29: "Col Q Y P"
    using P21 not_col_permutation_3 by blast
  have P30: "Q  C"
    using P15 P2 P4 between_equality_2 between_symmetry midpoint_bet by blast
  have P31: "Col Y B Z"
    using P28 by auto
  have P32: "Col Y Q' C'"
    by (simp add: P7 bet_col col_permutation_4)
  have P33: "Q  Q'"
    using P11 P15 P22 P25 P5 assms bet_neq12__neq col_transitivity_1 cong_reverse_identity by blast
  have P34: "C  C'"
    by (smt P15 P18 P3 P31 P8 assms bet_col col3 col_permutation_2 col_permutation_3 cong_3421 cong_diff midpoint_distinct_3)
  have P35: "Q Y C Z OFSC Q' Y C' Z"
    by (meson OFSC_def P15 P16 P2 P4 P5 P7 between_symmetry cong_3421 cong_reflexivity midpoint_bet not_cong_3412 outer_transitivity_between2)
  then have P36: "Cong C Z C' Z"
    using P19 five_segment_with_def by blast
  have P37: "Col Z Y X"
    by (simp add: P18 bet_col)
  have P38: "Y  Z"
    using P15 P3 cong_reverse_identity by blast
  then have P40: "X  Y"
    by (metis (mono_tags, opaque_lifting) Col_perm Cong_perm P14 P24 P25 P27 P36 P8 Per_def assms colx per_not_colp)
  have "Col A B X"
    using Col_perm P26 P31 P37 P38 col3 by blast
  thus ?thesis
    by (metis P18 P20 P27 P37 P40 Tarski_neutral_dimensionless.per_col Tarski_neutral_dimensionless_axioms assms between_equality col_permutation_3 l5_2 l8_16_2 l8_2)
qed

lemma l8_21_aux:
  assumes "¬ Col A B C"
  shows " P.  T. (A B Perp P A  Col A B T  Bet C T P)"
proof -
  obtain X where P1: "Col A B X  A B Perp C X"
    using assms l8_18_existence by blast
  have P2: "X PerpAt A B C X"
    by (simp add: P1 l8_15_1)
  have P3: "Per A X C"
    by (meson P1 Per_perm Tarski_neutral_dimensionless.l8_16_1 Tarski_neutral_dimensionless_axioms col_trivial_3)
  obtain C' where P4: "X Midpoint C C'  Cong A C A C'"
    using P3 Per_def by blast
  obtain C'' where P5: "A Midpoint C C''"
    using symmetric_point_construction by blast
  obtain P where P6: "P Midpoint C' C''"
    by (metis Cong_perm P4 P5 Tarski_neutral_dimensionless.Midpoint_def Tarski_neutral_dimensionless_axioms cong_inner_transitivity l7_25)
  have P7: "Per X A P"
    by (smt P3 P4 P5 P6 l7_2 l8_20_1_R2 l8_4 midpoint_distinct_3 symmetric_point_uniqueness)
  have P8: "X  C"
    using P1 assms by auto
  have P9: "A  P"
    using P4 P5 P6 P8 l7_9 midpoint_distinct_2 by blast
  obtain T where P10: "Bet P T C  Bet A T X"
    by (meson Mid_perm Midpoint_def P4 P5 P6 l3_17)
  have "A B Perp P A  Col A B T  Bet C T P"
  proof cases
    assume "A = X"
    thus ?thesis
      by (metis Bet_perm Col_def P1 P10 P9 between_identity col_trivial_3 perp_col2_bis)
  next
    assume "A  X"
    thus ?thesis
      by (metis Bet_perm Col_def P1 P10 P7 P9 Perp_perm col_transitivity_2 col_trivial_1 l8_3 per_perp perp_not_col2)
  qed
  thus ?thesis
    by blast
qed

lemma l8_21:
  assumes "A  B"
  shows " P T. A B Perp P A  Col A B T  Bet C T P"
  by (meson assms between_trivial2 l8_21_aux not_col_exists)

lemma per_cong:
  assumes "A  B" and
    "A  P" and
    "Per B A P" and
    "Per A B R" and
    "Cong A P B R" and
    "Col A B X" and
    "Bet P X R"
  shows "Cong A R P B"
proof -
  have P1: "Per P A B"
    using Per_cases assms(3) by blast
  obtain Q where P2: "R Midpoint B Q"
    using symmetric_point_construction by auto
  have P3: "B  R"
    using assms(2) assms(5) cong_identity by blast
  have P4: "Per A B Q"
    by (metis P2 P3 assms(1) assms(4) bet_neq23__neq col_permutation_4 midpoint_bet midpoint_col per_perp_in perp_in_col_perp_in perp_in_per)
  have P5: "Per P A X"
    using P1 assms(1) assms(6) per_col by blast
  have P6: "B  Q"
    using P2 P3 l7_3 by blast
  have P7: "Per R B X"
    by (metis assms(1) assms(4) assms(6) l8_2 not_col_permutation_4 per_col)
  have P8: "X  A"
    using P3 assms(1) assms(2) assms(3) assms(4) assms(7) bet_col per_not_colp by blast
  obtain P' where P9: "A Midpoint P P'"
    using Per_def assms(3) by blast
  obtain R' where P10: "Bet P' X R'  Cong X R' X R"
    using segment_construction by blast
  obtain M where P11: "M Midpoint R R'"
    by (meson P10 Tarski_neutral_dimensionless.l7_2 Tarski_neutral_dimensionless_axioms l7_25)
  have P12: "Per X M R"
    using P10 P11 Per_def cong_symmetry by blast
  have P13: "Cong X P X P'"
    using P9 assms(1) assms(3) assms(6) cong_left_commutativity l4_17 midpoint_cong per_double_cong by blast
  have P14: "X  P'"
    using P13 P8 P9 cong_identity l7_3 by blast
  have P15: "P  P'"
    using P9 assms(2) midpoint_distinct_2 by blast
  have P16: "¬ Col X P P'"
    using P13 P15 P8 P9 l7_17 l7_20 not_col_permutation_4 by blast
  have P17: "Bet A X M"
    using P10 P11 P13 P9 assms(7) cong_symmetry l7_22 by blast
  have P18: "X  R"
    using P3 P7 per_distinct_1 by blast
  have P19: "X  R'"
    using P10 P18 cong_diff_3 by blast
  have P20: "X  M"
    by (metis Col_def P10 P11 P16 P18 P19 assms(7) col_transitivity_1 midpoint_col)
  have P21: "M = B"
    by (smt Col_def P12 P17 P20 P8 Per_perm assms(1) assms(4) assms(6) col_transitivity_2 l8_3 l8_7)
  have "P X R P' OFSC P' X R' P"
    by (simp add: OFSC_def P10 P13 assms(7) cong_commutativity cong_pseudo_reflexivity cong_symmetry)
  then have "Cong R P' R' P"
    using P13 P14 cong_diff_3 five_segment_with_def by blast
  then have "P' A P R IFSC R' B R P"
    by (metis Bet_perm Cong_perm Midpoint_def P11 P21 P9 Tarski_neutral_dimensionless.IFSC_def Tarski_neutral_dimensionless_axioms assms(5) cong_mid2__cong cong_pseudo_reflexivity)
  thus ?thesis
    using l4_2 not_cong_1243 by blast
qed

lemma perp_cong:
  assumes "A  B" and
    "A  P" and
    "A B Perp P A" and
    "A B Perp R B" and
    "Cong A P B R" and
    "Col A B X" and
    "Bet P X R"
  shows "Cong A R P B"
  using Perp_cases assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) per_cong perp_per_1 by blast

lemma perp_exists:
  assumes "A  B"
  shows " X. PO X Perp A B"
proof cases
  assume "Col A B PO"
  then obtain C where P1: "A  C  B  C  PO  C  Col A B C"
    using diff_col_ex3 by blast
  then obtain P T where P2: "PO C Perp P PO  Col PO C T  Bet PO T P" using l8_21
    by blast
  then have "PO P Perp A B"
    by (metis P1 Perp_perm Col A B PO assms col3 col_trivial_2 col_trivial_3 perp_col2)
  thus ?thesis
    by blast
next
  assume "¬ Col A B PO"
  thus ?thesis using l8_18_existence
    using assms col_trivial_2 col_trivial_3 l8_18_existence perp_col0 by blast
qed

lemma perp_vector:
  assumes "A  B"
  shows " X Y. A B Perp X Y"
  using assms l8_21 by blast

lemma midpoint_existence_aux:
  assumes "A  B" and
    "A B Perp Q B" and
    "A B Perp P A" and
    "Col A B T" and
    "Bet Q T P" and
    "A P Le B Q"
  shows " X. X Midpoint A B"
proof -
  obtain R where P1: "Bet B R Q  Cong A P B R"
    using Le_def assms(6) by blast
  obtain X where P2: "Bet T X B  Bet R X P"
    using P1 assms(5) between_symmetry inner_pasch by blast
  have P3: "Col A B X"
    by (metis Col_def Out_cases P2 assms(4) between_equality l6_16_1 not_out_bet out_diff1)
  have P4: "B  R"
    using P1 assms(3) cong_identity perp_not_eq_2 by blast
  have P5: "¬ Col A B Q"
    using assms(2) col_trivial_2 l8_16_1 by blast
  have P6: "¬ Col A B R"
    using Col_def P1 P4 P5 l6_16_1 by blast
  have P7: "P  R"
    using P2 P3 P6 between_identity by blast
  have " X. X Midpoint A B"
  proof cases
    assume "A = P"
    thus ?thesis
      using assms(3) col_trivial_3 perp_not_col2 by blast
  next
    assume Q1: "¬ A = P"
    have Q2: "A B Perp R B"
      by (metis P1 P4 Perp_perm Tarski_neutral_dimensionless.bet_col1 Tarski_neutral_dimensionless_axioms assms(2) l5_1 perp_col1)
    then have Q3: "Cong A R P B"
      using P1 P2 P3 Q1 assms(1) assms(3) between_symmetry perp_cong by blast
    then have "X Midpoint A B  X Midpoint P R"
      by (smt P1 P2 P3 P6 P7 bet_col cong_left_commutativity cong_symmetry l7_2 l7_21 not_col_permutation_1)
    thus ?thesis
      by blast
  qed
  thus ?thesis by blast
qed

lemma midpoint_existence:
  " X. X Midpoint A B"
proof cases
  assume "A = B"
  thus ?thesis
    using l7_3_2 by blast
next
  assume P1: "¬ A = B"
  obtain Q where P2: "A B Perp B Q"
    by (metis P1 l8_21 perp_comm)
  obtain P T where P3: "A B Perp P A  Col A B T  Bet Q T P"
    using P2 l8_21_aux not_col_distincts perp_not_col2 by blast
  have P4: "A P Le B Q  B Q Le A P"
    by (simp add: local.le_cases)
  have P5: "A P Le B Q  ( X. X Midpoint A B)"
    by (meson P1 P2 P3 Tarski_neutral_dimensionless.Perp_cases Tarski_neutral_dimensionless.midpoint_existence_aux Tarski_neutral_dimensionless_axioms)
  have P6: "B Q Le A P  ( X. X Midpoint A B)"
  proof -
    {
      assume H1: "B Q Le A P"
      have Q6: "B  A"
        using P1 by auto
      have Q2: "B A Perp P A"
        by (simp add: P3 perp_left_comm)
      have Q3: "B A Perp Q B"
        using P2 Perp_perm by blast
      have Q4: "Col B A T"
        using Col_perm P3 by blast
      have Q5: "Bet P T Q"
        using Bet_perm P3 by blast
      obtain X where "X Midpoint B A"
        using H1 Q2 Q3 Q4 Q5 Q6 midpoint_existence_aux by blast
      then have " X. X Midpoint A B"
        using l7_2 by blast
    }
    thus ?thesis
      by simp
  qed
  thus ?thesis
    using P4 P5 by blast
qed

lemma perp_in_id:
  assumes "X PerpAt A B C A"
  shows "X = A"
  by (meson Col_cases assms col_trivial_3 l8_14_2_1b)

lemma l8_22:
  assumes "A  B" and
    "A  P" and
    "Per B A P" and
    "Per A B R" and
    "Cong A P B R" and
    "Col A B X" and
    "Bet P X R" and
    "Cong A R P B"
  shows "X Midpoint A B  X Midpoint P R"
  by (metis assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) assms(8) bet_col cong_commutativity cong_diff cong_right_commutativity l7_21 not_col_permutation_5 per_not_colp)

lemma l8_22_bis:
  assumes "A  B" and
    "A  P" and
    "A B Perp P A" and
    "A B Perp R B" and
    "Cong A P B R" and
    "Col A B X" and
    "Bet P X R"
  shows "Cong A R P B  X Midpoint A B  X Midpoint P R"
  by (metis l8_22 Perp_cases assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) perp_cong perp_per_2)

lemma perp_in_perp:
  assumes "X PerpAt A B C D"
  shows "A B Perp C D"
  using assms l8_14_2_1a by auto

lemma perp_proj:
  assumes "A B Perp C D" and
    "¬ Col A C D"
  shows " X. Col A B X  A X Perp C D"
  using assms(1) not_col_distincts by auto

lemma l8_24 :
  assumes "P A Perp A B" and
    "Q B Perp A B" and
    "Col A B T" and
    "Bet P T Q" and
    "Bet B R Q" and
    "Cong A P B R"
  shows " X. X Midpoint A B  X Midpoint P R"
proof -
  obtain X where P1: "Bet T X B  Bet R X P"
    using assms(4) assms(5) inner_pasch by blast
  have P2: "Col A B X"
    by (metis Out_cases P1 assms(3) bet_out_1 col_out2_col not_col_distincts out_trivial)
  have P3: "A  B"
    using assms(1) col_trivial_2 l8_16_1 by blast
  have P4: "A  P"
    using assms(1) col_trivial_1 l8_16_1 by blast
  have " X. X Midpoint A B  X Midpoint P R"
  proof cases
    assume "Col A B P"
    thus ?thesis
      using Perp_perm assms(1) perp_not_col by blast
  next
    assume Q1: "¬ Col A B P"
    have Q2: "B  R"
      using P4 assms(6) cong_diff by blast
    have Q3: "Q  B"
      using Q2 assms(5) between_identity by blast
    have Q4: "¬ Col A B Q"
      by (metis assms(2) col_permutation_3 l8_14_1 perp_col1 perp_not_col)
    have Q5: "¬ Col A B R"
      by (meson Q2 Q4 assms(5) bet_col col_transitivity_1 not_col_permutation_2)
    have Q6: "P  R"
      using P1 P2 Q5 between_identity by blast
    have " X. X Midpoint A B  X Midpoint P R"
    proof cases
      assume "A = P"
      thus ?thesis
        using P4 by blast
    next
      assume R0: "¬ A = P"
      have R1: "A B Perp R B"
        by (metis Perp_cases Q2 Tarski_neutral_dimensionless.bet_col1 Tarski_neutral_dimensionless_axioms assms(2) assms(5) bet_col col_transitivity_1 perp_col1)
      have R2: "Cong A R P B"
        using P1 P2 P3 Perp_perm R0 R1 assms(1) assms(6) between_symmetry perp_cong by blast
      have R3: "¬ Col A P B"
        using Col_perm Q1 by blast
      have R4: "P  R"
        by (simp add: Q6)
      have R5: "Cong A P B R"
        by (simp add: assms(6))
      have R6: "Cong P B R A"
        using R2 not_cong_4312 by blast
      have R7: "Col A X B"
        using Col_perm P2 by blast
      have R8: "Col P X R"
        by (simp add: P1 bet_col between_symmetry)
      thus ?thesis using l7_21
        using R3 R4 R5 R6 R7 by blast
    qed
    thus ?thesis by simp
  qed
  thus ?thesis
    by simp
qed

lemma col_per2__per:
  assumes "A  B" and
    "Col A B C" and
    "Per A X P" and
    "Per B X P"
  shows "Per C X P"
  by (meson Per_def assms(1) assms(2) assms(3) assms(4) l4_17 per_double_cong)

lemma perp_in_per_1:
  assumes "X PerpAt A B C D"
  shows "Per A X C"
  using PerpAt_def assms col_trivial_1 by auto

lemma perp_in_per_2:
  assumes "X PerpAt A B C D"
  shows "Per A X D"
  using assms perp_in_per_1 perp_in_right_comm by blast

lemma perp_in_per_3:
  assumes "X PerpAt A B C D"
  shows "Per B X C"
  using assms perp_in_comm perp_in_per_2 by blast

lemma perp_in_per_4:
  assumes "X PerpAt A B C D"
  shows "Per B X D"
  using assms perp_in_per_3 perp_in_right_comm by blast

subsection "Planes"

subsubsection "Coplanar"

lemma coplanar_perm_1:
  assumes "Coplanar A B C D"
  shows "Coplanar A B D C"
proof -
  obtain X where P1: "(Col A B X  Col C D X)  (Col A C X  Col B D X)  (Col A D X  Col B C X)"
    using Coplanar_def assms by blast
  then show ?thesis
    using Coplanar_def col_permutation_4 by blast
qed

lemma coplanar_perm_2:
  assumes "Coplanar A B C D"
  shows "Coplanar A C B D"
proof -
  obtain X where P1: "(Col A B X  Col C D X)  (Col A C X  Col B D X)  (Col A D X  Col B C X)"
    using Coplanar_def assms by blast
  then show ?thesis
    using Coplanar_def col_permutation_4 by blast
qed

lemma coplanar_perm_3:
  assumes "Coplanar A B C D"
  shows "Coplanar A C D B"
proof -
  obtain X where P1: "(Col A B X  Col C D X)  (Col A C X  Col B D X)  (Col A D X  Col B C X)"
    using Coplanar_def assms by blast
  then show ?thesis
    using Coplanar_def col_permutation_4 by blast
qed

lemma coplanar_perm_4:
  assumes "Coplanar A B C D"
  shows "Coplanar A D B C"
proof -
  obtain X where P1: "(Col A B X  Col C D X)  (Col A C X  Col B D X)  (Col A D X  Col B C X)"
    using Coplanar_def assms by blast
  then show ?thesis
    using Coplanar_def col_permutation_4 by blast
qed

lemma coplanar_perm_5:
  assumes "Coplanar A B C D"
  shows "Coplanar A D C B"
proof -
  obtain X where P1: "(Col A B X  Col C D X)  (Col A C X  Col B D X)  (Col A D X  Col B C X)"
    using Coplanar_def assms by blast
  then show ?thesis
    using Coplanar_def col_permutation_4 by blast
qed

lemma coplanar_perm_6:
  assumes "Coplanar A B C D"
  shows "Coplanar B A C D"
proof -
  obtain X where P1: "(Col A B X  Col C D X)  (Col A C X  Col B D X)  (Col A D X  Col B C X)"
    using Coplanar_def assms by blast
  then show ?thesis
    using Coplanar_def col_permutation_4 by blast
qed

lemma coplanar_perm_7:
  assumes "Coplanar A B C D"
  shows "Coplanar B A D C"
proof -
  obtain X where P1: "(Col A B X  Col C D X)  (Col A C X  Col B D X)  (Col A D X  Col B C X)"
    using Coplanar_def assms by blast
  then show ?thesis
    using Coplanar_def col_permutation_4 by blast
qed

lemma coplanar_perm_8:
  assumes "Coplanar A B C D"
  shows "Coplanar B C A D"
proof -
  obtain X where P1: "(Col A B X  Col C D X)  (Col A C X  Col B D X)  (Col A D X  Col B C X)"
    using Coplanar_def assms by blast
  then show ?thesis
    using Coplanar_def col_permutation_4 by blast
qed

lemma coplanar_perm_9:
  assumes "Coplanar A B C D"
  shows "Coplanar B C D A"
proof -
  obtain X where P1: "(Col A B X  Col C D X)  (Col A C X  Col B D X)  (Col A D X  Col B C X)"
    using Coplanar_def assms by blast
  then show ?thesis
    using Coplanar_def col_permutation_4 by blast
qed

lemma coplanar_perm_10:
  assumes "Coplanar A B C D"
  shows "Coplanar B D A C"
proof -
  obtain X where P1: "(Col A B X  Col C D X)  (Col A C X  Col B D X)  (Col A D X  Col B C X)"
    using Coplanar_def assms by blast
  then show ?thesis
    using Coplanar_def col_permutation_4 by blast
qed

lemma coplanar_perm_11:
  assumes "Coplanar A B C D"
  shows "Coplanar B D C A"
proof -
  obtain X where P1: "(Col A B X  Col C D X)  (Col A C X  Col B D X)  (Col A D X  Col B C X)"
    using Coplanar_def assms by blast
  then show ?thesis
    using Coplanar_def col_permutation_4 by blast
qed

lemma coplanar_perm_12:
  assumes "Coplanar A B C D"
  shows "Coplanar C A B D"
proof -
  obtain X where P1: "(Col A B X  Col C D X)  (Col A C X  Col B D X)  (Col A D X  Col B C X)"
    using Coplanar_def assms by blast
  then show ?thesis
    using Coplanar_def col_permutation_4 by blast
qed

lemma coplanar_perm_13:
  assumes "Coplanar A B C D"
  shows "Coplanar C A D B"
proof -
  obtain X where P1: "(Col A B X  Col C D X)  (Col A C X  Col B D X)  (Col A D X  Col B C X)"
    using Coplanar_def assms by blast
  then show ?thesis
    using Coplanar_def col_permutation_4 by blast
qed

lemma coplanar_perm_14:
  assumes "Coplanar A B C D"
  shows "Coplanar C B A D"
proof -
  obtain X where P1: "(Col A B X  Col C D X)  (Col A C X  Col B D X)  (Col A D X  Col B C X)"
    using Coplanar_def assms by blast
  then show ?thesis
    using Coplanar_def col_permutation_4 by blast
qed

lemma coplanar_perm_15:
  assumes "Coplanar A B C D"
  shows "Coplanar C B D A"
proof -
  obtain X where P1: "(Col A B X  Col C D X)  (Col A C X  Col B D X)  (Col A D X  Col B C X)"
    using Coplanar_def assms by blast
  then show ?thesis
    using Coplanar_def col_permutation_4 by blast
qed

lemma coplanar_perm_16:
  assumes "Coplanar A B C D"
  shows "Coplanar C D A B"
proof -
  obtain X where P1: "(Col A B X  Col C D X)  (Col A C X  Col B D X)  (Col A D X  Col B C X)"
    using Coplanar_def assms by blast
  then show ?thesis
    using Coplanar_def col_permutation_4 by blast
qed

lemma coplanar_perm_17:
  assumes "Coplanar A B C D"
  shows "Coplanar C D B A"
proof -
  obtain X where P1: "(Col A B X  Col C D X)  (Col A C X  Col B D X)  (Col A D X  Col B C X)"
    using Coplanar_def assms by blast
  then show ?thesis
    using Coplanar_def col_permutation_4 by blast
qed

lemma coplanar_perm_18:
  assumes "Coplanar A B C D"
  shows "Coplanar D A B C"
proof -
  obtain X where P1: "(Col A B X  Col C D X)  (Col A C X  Col B D X)  (Col A D X  Col B C X)"
    using Coplanar_def assms by blast
  then show ?thesis
    using Coplanar_def col_permutation_4 by blast
qed

lemma coplanar_perm_19:
  assumes "Coplanar A B C D"
  shows "Coplanar D A C B"
proof -
  obtain X where P1: "(Col A B X  Col C D X)  (Col A C X  Col B D X)  (Col A D X  Col B C X)"
    using Coplanar_def assms by blast
  then show ?thesis
    using Coplanar_def col_permutation_4 by blast
qed

lemma coplanar_perm_20:
  assumes "Coplanar A B C D"
  shows "Coplanar D B A C"
proof -
  obtain X where P1: "(Col A B X  Col C D X)  (Col A C X  Col B D X)  (Col A D X  Col B C X)"
    using Coplanar_def assms by blast
  then show ?thesis
    using Coplanar_def col_permutation_4 by blast
qed

lemma coplanar_perm_21:
  assumes "Coplanar A B C D"
  shows "Coplanar D B C A"
proof -
  obtain X where P1: "(Col A B X  Col C D X)  (Col A C X  Col B D X)  (Col A D X  Col B C X)"
    using Coplanar_def assms by blast
  then show ?thesis
    using Coplanar_def col_permutation_4 by blast
qed

lemma coplanar_perm_22:
  assumes "Coplanar A B C D"
  shows "Coplanar D C A B"
proof -
  obtain X where P1: "(Col A B X  Col C D X)  (Col A C X  Col B D X)  (Col A D X  Col B C X)"
    using Coplanar_def assms by blast
  then show ?thesis
    using Coplanar_def col_permutation_4 by blast
qed

lemma coplanar_perm_23:
  assumes "Coplanar A B C D"
  shows "Coplanar D C B A"
proof -
  obtain X where P1: "(Col A B X  Col C D X)  (Col A C X  Col B D X)  (Col A D X  Col B C X)"
    using Coplanar_def assms by blast
  then show ?thesis
    using Coplanar_def col_permutation_4 by blast
qed

lemma ncoplanar_perm_1:
  assumes "¬ Coplanar A B C D"
  shows "¬ Coplanar A B D C"
  using assms coplanar_perm_1 by blast

lemma ncoplanar_perm_2:
  assumes "¬ Coplanar A B C D"
  shows "¬ Coplanar A C B D"
  using assms coplanar_perm_2 by blast

lemma ncoplanar_perm_3:
  assumes "¬ Coplanar A B C D"
  shows "¬ Coplanar A C D B"
  using assms coplanar_perm_4 by blast

lemma ncoplanar_perm_4:
  assumes "¬ Coplanar A B C D"
  shows "¬ Coplanar A D B C"
  using assms coplanar_perm_3 by blast

lemma ncoplanar_perm_5:
  assumes "¬ Coplanar A B C D"
  shows "¬ Coplanar A D C B"
  using assms coplanar_perm_5 by blast

lemma ncoplanar_perm_6:
  assumes "¬ Coplanar A B C D"
  shows "¬ Coplanar B A C D"
  using assms coplanar_perm_6 by blast

lemma ncoplanar_perm_7:
  assumes "¬  Coplanar A B C D"
  shows "¬ Coplanar B A D C"
  using assms coplanar_perm_7 by blast

lemma ncoplanar_perm_8:
  assumes "¬ Coplanar A B C D"
  shows "¬ Coplanar B C A D"
  using assms coplanar_perm_12 by blast

lemma ncoplanar_perm_9:
  assumes "¬ Coplanar A B C D"
  shows "¬ Coplanar B C D A"
  using assms coplanar_perm_18 by blast

lemma ncoplanar_perm_10:
  assumes "¬ Coplanar A B C D"
  shows "¬ Coplanar B D A C"
  using assms coplanar_perm_13 by blast

lemma ncoplanar_perm_11:
  assumes "¬ Coplanar A B C D"
  shows "¬  Coplanar B D C A"
  using assms coplanar_perm_19 by blast

lemma ncoplanar_perm_12:
  assumes "¬ Coplanar A B C D"
  shows "¬ Coplanar C A B D"
  using assms coplanar_perm_8 by blast

lemma ncoplanar_perm_13:
  assumes "¬ Coplanar A B C D"
  shows "¬ Coplanar C A D B"
  using assms coplanar_perm_10 by blast

lemma ncoplanar_perm_14:
  assumes "¬ Coplanar A B C D"
  shows "¬ Coplanar C B A D"
  using assms coplanar_perm_14 by blast

lemma ncoplanar_perm_15:
  assumes "¬ Coplanar A B C D"
  shows "¬ Coplanar C B D A"
  using assms coplanar_perm_20 by blast

lemma ncoplanar_perm_16:
  assumes "¬ Coplanar A B C D"
  shows "¬ Coplanar C D A B"
  using assms coplanar_perm_16 by blast

lemma ncoplanar_perm_17:
  assumes "¬ Coplanar A B C D"
  shows "¬ Coplanar C D B A"
  using assms coplanar_perm_22 by blast

lemma ncoplanar_perm_18:
  assumes "¬ Coplanar A B C D"
  shows "¬ Coplanar D A B C"
  using assms coplanar_perm_9 by blast

lemma ncoplanar_perm_19:
  assumes "¬ Coplanar A B C D"
  shows "¬ Coplanar D A C B"
  using assms coplanar_perm_11 by blast

lemma ncoplanar_perm_20:
  assumes "¬ Coplanar A B C D"
  shows "¬ Coplanar D B A C"
  using assms coplanar_perm_15 by blast

lemma ncoplanar_perm_21:
  assumes "¬ Coplanar A B C D"
  shows "¬ Coplanar D B C A"
  using assms coplanar_perm_21 by blast

lemma ncoplanar_perm_22:
  assumes "¬ Coplanar A B C D"
  shows "¬ Coplanar D C A B"
  using assms coplanar_perm_17 by blast

lemma ncoplanar_perm_23:
  assumes "¬  Coplanar A B C D"
  shows "¬ Coplanar D C B A"
  using assms coplanar_perm_23 by blast

lemma coplanar_trivial:
  shows "Coplanar A A B C"
  using Coplanar_def NCol_cases col_trivial_1 by blast

lemma col__coplanar:
  assumes "Col A B C"
  shows "Coplanar A B C D"
  using Coplanar_def assms not_col_distincts by blast

lemma ncop__ncol:
  assumes "¬ Coplanar A B C D"
  shows "¬ Col A B C"
  using assms col__coplanar by blast

lemma ncop__ncols:
  assumes "¬ Coplanar A B C D"
  shows "¬ Col A B C  ¬ Col A B D  ¬ Col A C D  ¬ Col B C D"
  by (meson assms col__coplanar coplanar_perm_4 ncoplanar_perm_9)

lemma bet__coplanar:
  assumes "Bet A B C"
  shows "Coplanar A B C D"
  using assms bet_col ncop__ncol by blast

lemma out__coplanar:
  assumes "A Out B C"
  shows "Coplanar A B C D"
  using assms col__coplanar out_col by blast

lemma midpoint__coplanar:
  assumes "A Midpoint B C"
  shows "Coplanar A B C D"
  using assms midpoint_col ncop__ncol by blast

lemma perp__coplanar:
  assumes "A B Perp C D"
  shows "Coplanar A B C D"
proof -
  obtain P where "P PerpAt A B C D"
    using Perp_def assms by blast
  then show ?thesis
    using Coplanar_def perp_in_col by blast
qed

lemma ts__coplanar:
  assumes "A B TS C D"
  shows "Coplanar A B C D"
  by (metis (full_types) Coplanar_def TS_def assms bet_col col_permutation_2 col_permutation_3)

lemma reflectl__coplanar:
  assumes "A B ReflectL C D"
  shows "Coplanar A B C D"
  by (metis (no_types) ReflectL_def Tarski_neutral_dimensionless.perp__coplanar Tarski_neutral_dimensionless_axioms assms col__coplanar col_trivial_1 ncoplanar_perm_17)

lemma reflect__coplanar:
  assumes "A B Reflect C D"
  shows "Coplanar A B C D"
  by (metis (no_types) Reflect_def Tarski_neutral_dimensionless.reflectl__coplanar Tarski_neutral_dimensionless_axioms assms col_trivial_2 ncop__ncols)

lemma inangle__coplanar:
  assumes "A InAngle B C D"
  shows "Coplanar A B C D"
proof -
  obtain X where P1: "Bet B X D  (X = C  C Out X A)"
    using InAngle_def assms by auto
  then show ?thesis
    by (meson Col_cases Coplanar_def bet_col ncop__ncols out_col)
qed

lemma pars__coplanar:
  assumes "A B ParStrict C D"
  shows "Coplanar A B C D"
  using ParStrict_def assms by auto

lemma par__coplanar:
  assumes "A B Par C D"
  shows "Coplanar A B C D"
  using Par_def assms ncop__ncols pars__coplanar by blast

lemma plg__coplanar:
  assumes "Plg A B C D"
  shows "Coplanar A B C D"
proof -
  obtain M where "Bet A M C  Bet B M D"
    by (meson Plg_def assms midpoint_bet)
  then show ?thesis
    by (metis InAngle_def bet_out_1 inangle__coplanar ncop__ncols not_col_distincts)
qed

lemma plgs__coplanar:
  assumes "ParallelogramStrict A B C D"
  shows "Coplanar A B C D"
  using ParallelogramStrict_def assms par__coplanar by blast

lemma plgf__coplanar:
  assumes "ParallelogramFlat A B C D"
  shows "Coplanar A B C D"
  using ParallelogramFlat_def assms col__coplanar by auto

lemma parallelogram__coplanar:
  assumes "Parallelogram A B C D"
  shows "Coplanar A B C D"
  using Parallelogram_def assms plgf__coplanar plgs__coplanar by auto

lemma rhombus__coplanar:
  assumes "Rhombus A B C D"
  shows "Coplanar A B C D"
  using Rhombus_def assms plg__coplanar by blast

lemma rectangle__coplanar:
  assumes "Rectangle A B C D"
  shows "Coplanar A B C D"
  using Rectangle_def assms plg__coplanar by blast

lemma square__coplanar:
  assumes "Square A B C D"
  shows "Coplanar A B C D"
  using Square_def assms rectangle__coplanar by blast

lemma lambert__coplanar:
  assumes "Lambert A B C D"
  shows "Coplanar A B C D"
  using Lambert_def assms by presburger

subsubsection "Planes"

lemma ts_distincts:
  assumes "A B TS P Q"
  shows "A  B  A  P  A  Q  B  P  B  Q  P  Q"
  using TS_def assms bet_neq12__neq not_col_distincts by blast

lemma l9_2:
  assumes "A B TS P Q"
  shows "A B TS Q P"
  using TS_def assms between_symmetry by blast

lemma invert_two_sides:
  assumes "A B TS P Q"
  shows "B A TS P Q"
  using TS_def assms not_col_permutation_5 by blast

lemma l9_3:
  assumes "P Q TS A C" and
    "Col M P Q" and
    "M Midpoint A C" and
    "Col R P Q" and
    "R Out A B"
  shows "P Q TS B C"
proof -
  have P1: "¬ Col A P Q"
    using TS_def assms(1) by blast
  have P2: "P  Q"
    using P1 not_col_distincts by auto
  obtain T where P3: "Col T P Q  Bet A T C"
    using assms(2) assms(3) midpoint_bet by blast
  have P4: "A  C"
    using assms(1) ts_distincts by blast
  have P5: "T = M"
    by (smt P1 P3 Tarski_neutral_dimensionless.bet_col1 Tarski_neutral_dimensionless_axioms assms(2) assms(3) col_permutation_2 l6_21 midpoint_bet)
  have "P Q TS B C"
  proof cases
    assume "C = M"
    then show ?thesis
      using P4 assms(3) midpoint_distinct_1 by blast
  next
    assume P6: "¬ C = M"
    have P7: "¬ Col B P Q"
      by (metis P1 assms(4) assms(5) col_permutation_1 colx l6_3_1 out_col)
    have P97: "Bet R A B  Bet R B A"
      using Out_def assms(5) by auto
    {
      assume Q1: "Bet R A B"
      obtain B' where Q2: "M Midpoint B B'"
        using symmetric_point_construction by blast
      obtain R' where Q3: "M Midpoint R R'"
        using symmetric_point_construction by blast
      have Q4: "Bet B' C R'"
        using Q1 Q2 Q3 assms(3) between_symmetry l7_15 by blast
      obtain X where Q5: "Bet M X R'  Bet C X B"
        using Bet_perm Midpoint_def Q2 Q4 between_trivial2 l3_17 by blast
      have Q6: "Col X P Q"
      proof -
        have R1: "Col P M R"
          using P2 assms(2) assms(4) col_permutation_4 l6_16_1 by blast
        have R2: "Col Q M R"
          by (metis R1 assms(2) assms(4) col_permutation_5 l6_16_1 not_col_permutation_3)
        {
          assume "M = X"
          then have "Col X P Q"
            using assms(2) by blast
        }
        then have R3: "M = X  Col X P Q" by simp
        {
          assume "M  X"
          then have S1: "M  R'"
            using Q5 bet_neq12__neq by blast
          have "M  R"
            using Q3 S1 midpoint_distinct_1 by blast
          then have "Col X P Q"
            by (smt Col_perm Q3 Q5 R1 R2 S1 bet_out col_transitivity_2 midpoint_col out_col)
        }
        then have "M  X  Col X P Q" by simp
        then show ?thesis using R3 by blast
      qed
      have "Bet B X C"
        using Q5 between_symmetry by blast
      then have "P Q TS B C" using Q6
        using P7 TS_def assms(1) by blast
    }
    then have P98: "Bet R A B  P Q TS B C" by simp
    {
      assume S2: "Bet R B A"
      have S3: "Bet C M A"
        using Bet_perm P3 P5 by blast
      then obtain X where "Bet B X C  Bet M X R"
        using S2 inner_pasch by blast
      then have "P Q TS B C"
        by (metis Col_def P7 TS_def assms(1) assms(2) assms(4) between_inner_transitivity between_trivial l6_16_1 not_col_permutation_5)
    }
    then have "Bet R B A  P Q TS B C" by simp
    then show ?thesis using P97 P98
      by blast
  qed
  then show ?thesis by blast
qed

lemma mid_preserves_col:
  assumes "Col A B C" and
    "M Midpoint A A'" and
    "M Midpoint B B'" and
    "M Midpoint C C'"
  shows "Col A' B' C'"
  using Col_def assms(1) assms(2) assms(3) assms(4) l7_15 by auto

lemma per_mid_per:
  assumes (*"A ≠ B" and*)
    "Per X A B" and
    "M Midpoint A B" and
    "M Midpoint X Y"
  shows "Cong A X B Y  Per Y B A"
  by (meson Cong3_def Mid_perm assms(1) assms(2) assms(3) l7_13 l8_10)

lemma sym_preserve_diff:
  assumes "A  B" and
    "M Midpoint A A'" and
    "M Midpoint B B'"
  shows "A' B'"
  using assms(1) assms(2) assms(3) l7_9 by blast

lemma l9_4_1_aux_R1:
  assumes "R = S" and
    "S C Le R A" and
    "P Q TS A C" and
    "Col R P Q" and
    "P Q Perp A R" and
    "Col S P Q" and
    "P Q Perp C S" and
    "M Midpoint R S"
  shows " U C'. M Midpoint U C'  (R Out U A  S Out C C')"
proof -
  have P1: "M = R"
    using assms(1) assms(8) l7_3 by blast
  have P2: "¬ Col A P Q"
    using TS_def assms(3) by auto
  then have P3: "P  Q"
    using not_col_distincts by blast
  obtain T where P4: "Col T P Q  Bet A T C"
    using TS_def assms(3) by blast
  {
    assume "¬ M = T"
    then have "M PerpAt M T A M" using perp_col2
      by (metis P1 P4 assms(4) assms(5) not_col_permutation_3 perp_left_comm perp_perp_in)
    then have "M T Perp C M"
      using P1 P4 M  T assms(1) assms(4) assms(7) col_permutation_1 perp_col2 by blast
    then have "Per T M A"
      using M PerpAt M T A M perp_in_per_3 by blast
    have "Per T M C"
      by (simp add: M T Perp C M perp_per_1)
    have "M = T"
    proof -
      have "Per C M T"
        by (simp add: Per T M C l8_2)
      then show ?thesis using l8_6 l8_2
        using P4 Per T M A by blast
    qed
    then have "False"
      using M  T by blast
  }
  then have Q0: "M = T" by blast
  have R1: " U C'. ((M Midpoint U C'  M Out U A)  M Out C C')"
  proof -
    {
      fix U C'
      assume Q1: "M Midpoint U C'  M Out U A"
      have Q2: "C  M"
        using P1 assms(1) assms(7) perp_not_eq_2 by blast
      have Q3: "C'  M"
        using Q1 midpoint_not_midpoint out_diff1 by blast
      have Q4: "Bet U M C"
        using P4 Q0 Q1 bet_out__bet l6_6 by blast
      then have "M Out C C'"
        by (metis (full_types) Out_def Q1 Q2 Q3 l5_2 midpoint_bet)
    }
    then show ?thesis by blast
  qed
  have R2: " U C'. ((M Midpoint U C'  M Out C C')  M Out U A)"
  proof -
    {
      fix U C'
      assume Q1: "M Midpoint U C'  M Out C C'"
      have Q2: "C  M"
        using P1 assms(1) assms(7) perp_not_eq_2 by blast
      have Q3: "C'  M"
        using Q1 l6_3_1 by blast
      have Q4: "Bet U M C"
        by (metis Out_def Q1 between_inner_transitivity midpoint_bet outer_transitivity_between)
      then have "M Out U A"
        by (metis P2 P4 Q0 Q1 Q2 Q3 l6_2 midpoint_distinct_1)
    }
    then show ?thesis by blast
  qed
  then show ?thesis
    using R1 P1 P2 assms by blast
qed

lemma l9_4_1_aux_R21:
  assumes "R  S" and
    "S C Le R A" and
    "P Q TS A C" and
    "Col R P Q" and
    "P Q Perp A R" and
    "Col S P Q" and
    "P Q Perp C S" and
    "M Midpoint R S"
  shows " U C'. M Midpoint U C'   (R Out U A  S Out C C')"
proof -
  obtain D where P1: "Bet R D A  Cong S C R D"
    using Le_def assms(2) by blast
  have P2: "C  S"
    using assms(7) perp_not_eq_2 by auto
  have P3: "R  D"
    using P1 P2 cong_identity by blast
  have P4: "R S Perp A R"
    using assms(1) assms(4) assms(5) assms(6) not_col_permutation_2 perp_col2 by blast
  have " M. (M Midpoint S R  M Midpoint C D)"
  proof -
    have Q1: "¬ Col A P Q"
      using TS_def assms(3) by blast
    have Q2: "P  Q"
      using Q1 not_col_distincts by blast
    obtain T where Q3: "Col T P Q  Bet A T C"
      using TS_def assms(3) by blast
    have Q4: "C S Perp S R"
      by (metis NCol_perm assms(1) assms(4) assms(6) assms(7) perp_col0)
    have Q5: "A R Perp S R"
      using P4 Perp_perm by blast
    have Q6: "Col S R T"
      using Col_cases Q2 Q3 assms(4) assms(6) col3 by blast
    have Q7: "Bet C T A"
      using Bet_perm Q3 by blast
    have Q8: "Bet R D A"
      by (simp add: P1)
    have "Cong S C R D"
      by (simp add: P1)
    then show ?thesis using P1 Q4 Q5 Q6 Q7 l8_24 by blast
  qed
  then obtain M' where P5: "M' Midpoint S R  M' Midpoint C D" by blast
  have P6: "M = M'"
    by (meson P5 assms(8) l7_17_bis)
  have L1: " U C'. (M Midpoint U C'  R Out U A)  S Out C C'"
  proof -
    {
      fix U C'
      assume R1: "M Midpoint U C'  R Out U A"
      have R2: "C  S"
        using P2 by auto
      have R3: "C'  S"
        using P5 R1 P6 l7_9_bis out_diff1 by blast
      have R4: "Bet S C C'  Bet S C' C"
      proof -
        have R5: "Bet R U A  Bet R A U"
          using Out_def R1 by auto
        {
          assume "Bet R U A"
          then have "Bet R U D  Bet R D U"
            using P1 l5_3 by blast
          then have "Bet S C C'  Bet S C' C"
            using P5 P6 R1 l7_15 l7_2 by blast
        }
        then have R6: "Bet R U A  Bet S C C'  Bet S C' C" by simp
        have "Bet R A U  Bet S C C'  Bet S C' C"
          using P1 P5 P6 R1 between_exchange4 l7_15 l7_2 by blast
        then show ?thesis using R5 R6 by blast
      qed
      then have "S Out C C'"
        by (simp add: Out_def R2 R3)
    }
    then show ?thesis by simp
  qed
  have " U C'. (M Midpoint U C'  S Out C C')  R Out U A"
  proof -
    {
      fix U C'
      assume Q1: "M Midpoint U C'  S Out C C'"
      then have Q2: "U  R"
        using P5 P6 l7_9_bis out_diff2 by blast
      have Q3: "A  R"
        using assms(5) perp_not_eq_2 by auto
      have Q4: "Bet S C C'  Bet S C' C"
        using Out_def Q1 by auto
      {
        assume V0: "Bet S C C'"
        have V1: "R  D"
          by (simp add: P3)
        then have V2: "Bet R D U"
        proof -
          have W1: "M Midpoint S R"
            using P5 P6 by blast
          have W2: "M Midpoint C D"
            by (simp add: P5 P6)
          have "M Midpoint C' U"
            by (simp add: Q1 l7_2)
          then show ?thesis
            using V0 P5 P6 l7_15 by blast
        qed
        have "Bet R D A"
          using P1 by auto
        then have "Bet R U A  Bet R A U"
          using V1 V2 l5_1 by blast
      }
      then have Q5: "Bet S C C'  Bet R U A  Bet R A U" by simp
      {
        assume R1: "Bet S C' C"
        have "Bet R U A"
          using P1 P5 P6 Q1 R1 between_exchange4 l7_15 l7_2 by blast
      }
      then have "Bet S C' C  Bet R U A  Bet R A U" by simp
      then have "Bet R U A  Bet R A U"
        using Q4 Q5 by blast
      then have "R Out U A"
        by (simp add: Out_def Q2 Q3)
    }
    then show ?thesis by simp
  qed
  then show ?thesis
    using L1 by blast
qed

lemma l9_4_1_aux:
  assumes "S C Le R A" and
    "P Q TS A C" and
    "Col R P Q" and
    "P Q Perp A R" and
    "Col S P Q" and
    "P Q Perp C S" and
    "M Midpoint R S"
  shows " U C'. (M Midpoint U C'  (R Out U A  S Out C C'))"
  using l9_4_1_aux_R1 l9_4_1_aux_R21 assms by smt

lemma per_col_eq:
  assumes "Per A B C" and
    "Col A B C" and
    "B  C"
  shows "A = B"
  using assms(1) assms(2) assms(3) l8_9 by blast

lemma l9_4_1:
  assumes "P Q TS A C" and
    "Col R P Q" and
    "P Q Perp A R" and
    "Col S P Q" and
    "P Q Perp C S" and
    "M Midpoint R S"
  shows " U C'. M Midpoint U C'  (R Out U A  S Out C C')"
proof -
  have P1: "S C Le R A  R A  Le S C"
    using local.le_cases by blast
  {
    assume Q1: "S C Le R A"
    {
      fix U C'
      assume "M Midpoint U C'"
      then have "(R Out U A  S Out C C')"
        using Q1 assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) l9_4_1_aux by blast
    }
    then have " U C'. M Midpoint U C'  (R Out U A  S Out C C')" by simp
  }
  then have P2: "S C Le R A  ( U C'. M Midpoint U C'  (R Out U A  S Out C C'))" by simp
  {
    assume Q2: " R A Le S C"
    {
      fix U C'
      assume "M Midpoint U C'"
      then have "(R Out A U  S Out C' C)"
        using Q2 assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) l7_2 l9_2 l9_4_1_aux by blast
      then have "(R Out U A  S Out C C')"
        using l6_6 by blast
    }
    then have " U C'. M Midpoint U C'  (R Out U A  S Out C C')" by simp
  }
  then have P3: "R A Le S C  ( U C'. M Midpoint U C'  (R Out U A  S Out C C'))" by simp

  then show ?thesis
    using P1 P2 by blast
qed

lemma mid_two_sides:
  assumes "M Midpoint A B" and
    "¬ Col A B X" and
    "M Midpoint X Y"
  shows "A B TS X Y"
proof -
  have f1: "¬ Col Y A B"
    by (meson Mid_cases Tarski_neutral_dimensionless.mid_preserves_col Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) col_permutation_3)
  have "Bet X M Y"
    using assms(3) midpoint_bet by blast
  then show ?thesis
    using f1 by (metis (no_types) TS_def assms(1) assms(2) col_permutation_1 midpoint_col)
qed


lemma col_preserves_two_sides:
  assumes "C  D" and
    "Col A B C" and
    "Col A B D" and
    "A B TS X Y"
  shows "C D TS X Y"
proof -
  have P1: "¬ Col X A B"
    using TS_def assms(4) by blast
  then have P2: "A  B"
    using not_col_distincts by blast
  have P3: "¬ Col X C D"
    by (metis Col_cases P1 Tarski_neutral_dimensionless.colx Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3))
  have P4: "¬ Col Y C D"
    by (metis Col_cases TS_def Tarski_neutral_dimensionless.colx Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) assms(4))
  then show ?thesis
  proof -
    obtain pp :: "'p  'p  'p  'p  'p" where
      "x0 x1 x2 x3. (v4. Col v4 x3 x2  Bet x1 v4 x0) = (Col (pp x0 x1 x2 x3) x3 x2  Bet x1 (pp x0 x1 x2 x3) x0)"
      by moura
    then have f1: "¬ Col X A B  ¬ Col Y A B  Col (pp Y X B A) A B  Bet X (pp Y X B A) Y"
      using TS_def assms(4) by presburger
    then have "Col (pp Y X B A) C D"
      by (meson P2 assms(2) assms(3) col3 not_col_permutation_3 not_col_permutation_4)
    then show ?thesis
      using f1 TS_def P3 P4 by blast
  qed
qed

lemma out_out_two_sides:
  assumes "A  B" and
    "A B TS X Y" and
    "Col I A B" and
    "Col I X Y" and
    "I Out X U" and
    "I Out Y V"
  shows "A B TS U V"
proof -
  have P0: "¬ Col X A B"
    using TS_def assms(2) by blast
  then have P1: "¬ Col V A B"
    by (smt assms(2) assms(3) assms(4) assms(6) col_out2_col col_transitivity_1 not_col_permutation_3 not_col_permutation_4 out_diff2 out_trivial ts_distincts)
  have P2: "¬ Col U A B"
    by (metis P0 assms(3) assms(5) col_permutation_2 colx out_col out_distinct)
  obtain T where P3: "Col T A B  Bet X T Y"
    using TS_def assms(2) by blast
  have "I = T"
  proof -
    have f1: "p pa pb. ¬ Col p pa pb  ¬ Col p pb pa  ¬ Col pa p pb  ¬ Col pa pb p  ¬ Col pb p pa  ¬ Col pb pa p  Col p pa pb"
      using Col_cases by blast
    then have f2: "Col X Y I"
      using assms(4) by blast
    have f3: "Col B A I"
      using f1 assms(3) by blast
    have f4: "Col B A T"
      using f1 P3 by blast
    have f5: "¬ Col X A B  ¬ Col X B A  ¬ Col A X B  ¬ Col A B X  ¬ Col B X A  ¬ Col B A X"
      using f1 ¬ Col X A B by blast
    have f6: "A  B  A  X  A  Y  B  X  B  Y  X  Y"
      using assms(2) ts_distincts by presburger
    have "Col X Y T"
      using f1 by (meson P3 bet_col)
    then show ?thesis
      using f6 f5 f4 f3 f2 by (meson Tarski_neutral_dimensionless.l6_21 Tarski_neutral_dimensionless_axioms)
  qed
  then have "Bet U T V"
    using P3 assms(5) assms(6) bet_out_out_bet by blast
  then show ?thesis
    using P1 P2 P3 TS_def by blast
qed

lemma l9_4_2_aux_R1:
  assumes "R = S " and
    "S C Le R A" and
    "P Q TS A C" and
    "Col R P Q" and
    "P Q Perp A R" and
    "Col S P Q" and
    "P Q Perp C S" and
    "R Out U A" and
    "S Out V C"
  shows "P Q TS U V"
proof -
  have "¬ Col A P Q"
    using TS_def assms(3) by auto
  then have P2: "P  Q"
    using not_col_distincts by blast
  obtain T where P3: "Col T P Q  Bet A T C"
    using TS_def assms(3) by blast
  have "R = T" using assms(1) assms(5) assms(6) assms(7) col_permutation_1 l8_16_1 l8_6
    by (meson P3)
  then show ?thesis
    by (smt P2 P3 assms(1) assms(3) assms(8) assms(9) bet_col col_transitivity_2 l6_6 not_col_distincts out_out_two_sides)
qed

lemma l9_4_2_aux_R2:
  assumes "R  S" and
    "S C Le R A" and
    "P Q TS A C" and
    "Col R P Q" and
    "P Q Perp A R" and
    "Col S P Q" and
    "P Q Perp C S" and
    "R Out U A" and
    "S Out V C"
  shows "P Q TS U V"
proof -
  have P1: "P  Q"
    using assms(7) perp_distinct by auto
  have P2: "R S TS A C"
    using assms(1) assms(3) assms(4) assms(6) col_permutation_1 col_preserves_two_sides by blast
  have P3: "Col R S P"
    using P1 assms(4) assms(6) col2__eq not_col_permutation_1 by blast
  have P4: "Col R S Q"
    by (metis P3 Tarski_neutral_dimensionless.colx Tarski_neutral_dimensionless_axioms assms(4) assms(6) col_trivial_2)
  have P5: "R S Perp A R"
    using NCol_perm assms(1) assms(4) assms(5) assms(6) perp_col2 by blast
  have P6: "R S Perp C S"
    using assms(1) assms(4) assms(6) assms(7) col_permutation_1 perp_col2 by blast
  have P7: "¬ Col A R S"
    using P2 TS_def by blast
  obtain T where P8: "Col T R S  Bet A T C"
    using P2 TS_def by blast
  obtain C' where P9: "Bet R C' A  Cong S C R C'"
    using Le_def assms(2) by blast
  have " X. X Midpoint S R  X Midpoint C C'"
  proof -
    have Q1: "C S Perp S R"
      using P6 Perp_perm by blast
    have Q2: "A R Perp S R"
      using P5 Perp_perm by blast
    have Q3: "Col S R T"
      using Col_cases P8 by blast
    have Q4: "Bet C T A"
      using Bet_perm P8 by blast
    have Q5: "Bet R C' A"
      by (simp add: P9)
    have "Cong S C R C'"
      by (simp add: P9)
    then show ?thesis using Q1 Q2 Q3 Q4 Q5 l8_24
      by blast
  qed
  then obtain M where P10: "M Midpoint S R  M Midpoint C C'" by blast
  obtain U' where P11: "M Midpoint U U'"
    using symmetric_point_construction by blast
  have P12: "R  U"
    using assms(8) out_diff1 by blast
  have P13: "R S TS U U'"
    by (smt P10 P11 P12 P7 assms(8) col_transitivity_2 invert_two_sides mid_two_sides not_col_permutation_3 not_col_permutation_4 out_col)
  have P14: "R S TS V U"
  proof -
    have Q1: "Col M R S"
      using P10 midpoint_col not_col_permutation_5 by blast
    have Q2: "M Midpoint U' U"
      by (meson P11 Tarski_neutral_dimensionless.Mid_cases Tarski_neutral_dimensionless_axioms)
    have "S Out U' V"
      by (meson P10 P11 P2 P5 P6 Tarski_neutral_dimensionless.l7_2 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(8) assms(9) l6_6 l6_7 l9_4_1_aux_R21 not_col_distincts)
    then show ?thesis
      using P13 Q1 Q2 col_trivial_3 l9_2 l9_3 by blast
  qed
  then show ?thesis
    using P1 P3 P4 col_preserves_two_sides l9_2 by blast
qed

lemma l9_4_2_aux:
  assumes "S C Le R A" and
    "P Q TS A C" and
    "Col R P Q" and
    "P Q Perp A R" and
    "Col S P Q" and
    "P Q Perp C S" and
    "R Out U A" and
    "S Out V C"
  shows "P Q TS U V"
  using l9_4_2_aux_R1 l9_4_2_aux_R2
  by (metis assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) assms(8))

lemma l9_4_2:
  assumes "P Q TS A C" and
    "Col R P Q" and
    "P Q Perp A R" and
    "Col S P Q" and
    "P Q Perp C S" and
    "R Out U A" and
    "S Out V C"
  shows "P Q TS U V"
proof -
  have P1: "S C Le R A  R A Le S C"
    by (simp add: local.le_cases)
  have P2: "S C Le R A  P Q TS U V"
    by (simp add: assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) l9_4_2_aux)
  have "R A Le S C  P Q TS U V"
    by (simp add: assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) l9_2 l9_4_2_aux)
  then show ?thesis
    using P1 P2 by blast
qed

lemma l9_5:
  assumes "P Q TS A C" and
    "Col R P Q" and
    "R Out A B"
  shows "P Q TS B C"
proof -
  have P1: "P  Q"
    using assms(1) ts_distincts by blast
  obtain A' where P2: "Col P Q A'  P Q Perp A A'"
    by (metis NCol_perm Tarski_neutral_dimensionless.TS_def Tarski_neutral_dimensionless_axioms assms(1) l8_18_existence)
  obtain C' where P3: "Col P Q C'  P Q Perp C C'"
    using Col_perm TS_def assms(1) l8_18_existence by blast
  obtain M where P5: "M Midpoint A' C'"
    using midpoint_existence by blast
  obtain D where S2: "M Midpoint A D"
    using symmetric_point_construction by auto
  have " B0. Col P Q B0  P Q Perp B B0"
  proof -
    have S1: "¬ Col P Q B"
      by (metis P2 Tarski_neutral_dimensionless.colx Tarski_neutral_dimensionless.perp_not_col2 Tarski_neutral_dimensionless_axioms assms(2) assms(3) col_permutation_1 l6_3_1 out_col)
    then show ?thesis
      by (simp add: l8_18_existence)
  qed
  then obtain B' where P99: "Col P Q B'  P Q Perp B B'" by blast
  have "P Q TS B C"
  proof -
    have S3: "C' Out D C  A' Out A A"
      using Out_cases P2 P3 P5 S2 assms(1) l9_4_1 not_col_permutation_1 by blast
    then have S4: "C' Out D C"
      using P2 Tarski_neutral_dimensionless.perp_not_eq_2 Tarski_neutral_dimensionless_axioms out_trivial by fastforce
    have S5: "P Q TS A D"
      using P2 P3 S3 S4 assms(1) col_permutation_2 l9_4_2 by blast
    {
      assume "A'  C'"
      then have "Col M P Q"
        by (smt P2 P3 P5 col_trivial_2 l6_21 midpoint_col not_col_permutation_1)
      then have "P Q TS B D"
        using S2 S5 assms(2) assms(3) l9_3 by blast
    }
    then have "A'  C'  P Q TS B D" by simp
    then have S6: "P Q TS B D"
      by (metis P3 P5 S2 S5 assms(2) assms(3) l9_3 midpoint_distinct_2 not_col_permutation_1)
    have S7: "Col B' P Q"
      using Col_perm P99 by blast
    have S8: "P Q Perp B B'"
      using P99 by blast
    have S9: "Col C' P Q"
      using Col_cases P3 by auto
    have S10: "P Q Perp D C'"
      by (metis Col_perm P3 S4 l6_3_1 out_col perp_col1 perp_right_comm)
    have S11: "B' Out B B"
      by (metis (no_types) P99 out_trivial perp_not_eq_2)
    have "C' Out C D"
      by (simp add: S4 l6_6)
    then show ?thesis using S6 S7 S8 S9 S10 S11 l9_4_2 by blast
  qed
  then show ?thesis using l8_18_existence by blast
qed

lemma outer_pasch_R1:
  assumes "Col P Q C" and
    "Bet A C P" and
    "Bet B Q C"
  shows " X. Bet A X B  Bet P Q X"
  by (smt Bet_perm Col_def assms(1) assms(2) assms(3) between_exchange3 between_trivial outer_transitivity_between2)

lemma outer_pasch_R2:
  assumes "¬ Col P Q C" and
    "Bet A C P" and
    "Bet B Q C"
  shows " X. Bet A X B  Bet P Q X"
proof cases
  assume "B = Q"
  then show ?thesis
    using between_trivial by blast
next
  assume P1: "B  Q"
  have P2: "A  P"
    using assms(1) assms(2) between_identity col_trivial_3 by blast
  have P3: "P  Q"
    using assms(1) col_trivial_1 by blast
  have P4: "P  B"
    using assms(1) assms(3) bet_col by blast
  have P5: "P Q TS C B"
  proof -
    have Q1: "¬ Col C P Q"
      using Col_cases assms(1) by blast
    have Q2: "¬ Col B P Q"
      by (metis Col_cases P1 Tarski_neutral_dimensionless.colx Tarski_neutral_dimensionless_axioms assms(1) assms(3) bet_col col_trivial_2)
    have " T. Col T P Q  Bet C T B"
      using Col_cases assms(3) between_symmetry col_trivial_2 by blast
    then show ?thesis
      by (simp add: Q1 Q2 TS_def)
  qed
  have P6: "P Q TS A B"
    by (metis P5 assms(1) assms(2) bet_out_1 l9_5 not_col_distincts)
  obtain X where P7: "Col X P Q  Bet A X B"
    using P6 TS_def by blast
  have "Bet P Q X"
  proof -
    obtain T where P8: "Bet X T P  Bet C T B"
      using P7 assms(2) between_symmetry inner_pasch by blast
    have P9: "B  C"
      using P1 assms(3) bet_neq12__neq by blast
    have P10: "T = Q"
    proof -
      have f1: "p pa pb. Col pb pa p  ¬ Bet pb pa p"
        by (meson bet_col1 between_trivial)
      then have f2: "Col Q C B"
        using NCol_cases assms(3) by blast
      have "Col T C B"
        using f1 NCol_cases P8 by blast
      then show ?thesis
        using f2 f1 by (metis (no_types) NCol_cases P7 P8 assms(1) between_trivial l6_16_1 l6_2 not_bet_and_out)
    qed
    then show ?thesis
      using P8 between_symmetry by blast
  qed
  then show ?thesis using P7 by blast
qed

lemma outer_pasch:
  assumes "Bet A C P" and
    "Bet B Q C"
  shows " X. Bet A X B  Bet P Q X"
  using assms(1) assms(2) outer_pasch_R1 outer_pasch_R2 by blast

lemma os_distincts:
  assumes "A B OS X Y"
  shows "A  B  A  X  A  Y  B  X  B  Y"
  using OS_def assms ts_distincts by blast

lemma invert_one_side:
  assumes "A B OS P Q"
  shows "B A OS P Q"
proof -
  obtain T where "A B TS P T  A B TS Q T"
    using OS_def assms by blast
  then have "B A TS P T  B A TS Q T"
    using invert_two_sides by blast
  thus ?thesis
    using OS_def by blast
qed

lemma l9_8_1:
  assumes "P Q TS A C" and
    "P Q TS B C"
  shows "P Q OS A B"
proof -
  have " R::'p. (P Q TS A R  P Q TS B R)"
    using assms(1) assms(2) by blast
  then show ?thesis
    using OS_def by blast
qed

lemma not_two_sides_id:
  shows "¬ P Q TS A A"
  using ts_distincts by blast

lemma l9_8_2:
  assumes "P Q TS A C" and
    "P Q OS A B"
  shows "P Q TS B C"
proof -
  obtain D where P1: "P Q TS A D  P Q TS B D"
    using assms(2) OS_def by blast
  then have "P  Q"
    using ts_distincts by blast
  obtain T where P2: "Col T P Q  Bet A T C"
    using TS_def assms(1) by blast
  obtain X where P3: "Col X P Q  Bet A X D"
    using TS_def P1 by blast
  obtain Y where P4: "Col Y P Q  Bet B Y D"
    using TS_def P1 by blast
  then obtain M where P5: "Bet Y M A  Bet X M B" using P3 inner_pasch by blast
  have P6: "A  D"
    using P1 ts_distincts by blast
  have P7: "B  D"
    using P1 not_two_sides_id by blast
  {
    assume Q0: "Col A B D"
    have "P Q TS B C"
    proof cases
      assume Q1: "M = Y"
      have "X = Y"
      proof -
        have S1: "¬ Col P Q A"
          using TS_def assms(1) not_col_permutation_1 by blast
        have S3: "Col P Q X"
          using Col_perm P3 by blast
        have S4: "Col P Q Y"
          using Col_perm P4 by blast
        have S5: "Col A D X"
          by (simp add: P3 bet_col col_permutation_5)
        have "Col A D Y"
          by (metis Col_def P5 Q1 S5 Q0 between_equality between_trivial l6_16_1)
        then show ?thesis using S1 S3 S4 S5 P6 l6_21
          by blast
      qed
      then have "X Out A B"
        by (metis P1 P3 P4 TS_def l6_2)
      then show ?thesis using assms(1) P3 l9_5 by blast
    next
      assume Z1: "¬ M = Y"
      have "X = Y"
      proof -
        have S1: "¬ Col P Q A"
          using TS_def assms(1) not_col_permutation_1 by blast
        have S3: "Col P Q X"
          using Col_perm P3 by blast
        have S4: "Col P Q Y"
          using Col_perm P4 by blast
        have S5: "Col A D X"
          by (simp add: P3 bet_col col_permutation_5)
        have "Col A D Y"
          by (metis Col_def P4 Q0 P7 l6_16_1)
        then show ?thesis using S1 S3 S4 S5 P6 l6_21
          by blast
      qed
      then have Z3: "M  X" using Z1 by blast
      have Z4: "P Q TS M C"
        by (meson Out_cases P4 P5 Tarski_neutral_dimensionless.l9_5 Tarski_neutral_dimensionless_axioms Z1 assms(1) bet_out)
      have "X Out M B"
        using P5 Z3 bet_out by auto
      then show ?thesis using Z4 P3 l9_5 by blast
    qed
  }
  then have Z99: "Col A B D  P Q TS B C" by blast
  {
    assume Q0: "¬ Col A B D"
    have Q1: "P Q TS M C"
    proof -
      have S3: "Y Out A M"
      proof -
        have T1: "A  Y"
          using Col_def P4 Q0 col_permutation_4 by blast
        have T2: "M  Y"
        proof -
          {
            assume T3: "M = Y"
            have "Col B D X"
            proof -
              have U1: "B  M"
                using P1 P4 T3 TS_def by blast
              have U2: "Col B M D"
                by (simp add: P4 T3 bet_col)
              have "Col B M X"
                by (simp add: P5 bet_col between_symmetry)
              then show ?thesis using U1 U2
                using col_transitivity_1 by blast
            qed
            have "False"
              by (metis NCol_cases P1 P3 TS_def Col B D X Q0 bet_col col_trivial_2 l6_21)
          }
          then show ?thesis by blast
        qed
        have "Bet Y A M  Bet Y M A" using P5 by blast
        then show ?thesis using T1 T2
          by (simp add: Out_def)
      qed
      then have "X Out M B"
        by (metis P1 P3 P4 P5 TS_def bet_out l9_5)
      then show ?thesis using assms(1) S3 l9_5 P3 P4 by blast
    qed
    have "X Out M B"
      by (metis P3 P5 Q1 TS_def bet_out)
    then have "P Q TS B C" using Q1 P3 l9_5 by blast
  }
  then have "¬ Col A B D  P Q TS B C" by blast
  then show ?thesis using Z99 by blast
qed

lemma l9_9:
  assumes "P Q TS A B"
  shows "¬ P Q OS A B"
  using assms l9_8_2 not_two_sides_id by blast

lemma l9_9_bis:
  assumes "P Q OS A B"
  shows "¬ P Q TS A B"
  using assms l9_9 by blast

lemma one_side_chara:
  assumes "P Q OS A B"
  shows " X. Col X P Q  ¬ Bet A X B"
proof -
  have "¬ Col A P Q  ¬ Col B P Q"
    using OS_def TS_def assms by auto
  then show ?thesis
    using l9_9_bis TS_def assms by blast
qed

lemma l9_10:
  assumes "¬ Col A P Q"
  shows " C. P Q TS A C"
  by (meson Col_perm assms mid_two_sides midpoint_existence symmetric_point_construction)

lemma one_side_reflexivity:
  assumes "¬ Col A P Q"
  shows "P Q OS A A"
  using assms l9_10 l9_8_1 by blast

lemma one_side_symmetry:
  assumes "P Q OS A B"
  shows "P Q OS B A"
  by (meson Tarski_neutral_dimensionless.OS_def Tarski_neutral_dimensionless_axioms assms invert_two_sides)

lemma one_side_transitivity:
  assumes "P Q OS A B" and
    "P Q OS B C"
  shows "P Q OS A C"
  by (meson Tarski_neutral_dimensionless.OS_def Tarski_neutral_dimensionless.l9_8_2 Tarski_neutral_dimensionless_axioms assms(1) assms(2))

lemma l9_17:
  assumes "P Q OS A C" and
    "Bet A B C"
  shows "P Q OS A B"
proof cases
  assume "A = C"
  then show ?thesis
    using assms(1) assms(2) between_identity by blast
next
  assume P1: "¬ A = C"
  obtain D where P2: "P Q TS A D  P Q TS C D"
    using OS_def assms(1) by blast
  then have P3: "P  Q"
    using ts_distincts by blast
  obtain X where P4: "Col X P Q  Bet A X D"
    using P2 TS_def by blast
  obtain Y where P5: "Col Y P Q  Bet C Y D"
    using P2 TS_def by blast
  obtain T where P6: "Bet B T D  Bet X T Y"
    using P4 P5 assms(2) l3_17 by blast
  have P7: "P Q TS A D"
    by (simp add: P2)
  have "P Q TS B D"
  proof -
    have Q1: "¬ Col B P Q"
      using assms(1) assms(2) one_side_chara by blast
    have Q2: "¬ Col D P Q"
      using P2 TS_def by blast
    obtain T0 where "Col T0 P Q  Bet B T0 D"
    proof -
      assume a1: "T0. Col T0 P Q  Bet B T0 D  thesis"
      obtain pp :: 'p where
        f2: "Bet B pp D  Bet X pp Y"
        using thesis. (T. Bet B T D  Bet X T Y  thesis)  thesis by blast
      have "Col P Q Y"
        using Col_def P5 by blast
      then have "Y = X  Col P Q pp"
        using f2 Col_def P4 colx by blast
      then show ?thesis
        using f2 a1 by (metis BetSEq BetS_def Col_def P4)
    qed
    then show ?thesis using Q1 Q2
      using TS_def by blast
  qed
  then show ?thesis using P7
    using OS_def by blast
qed

lemma l9_18_R1:
  assumes "Col X Y P" and
    "Col A B P"
    and "X Y TS A B"
  shows "Bet A P B  ¬ Col X Y A  ¬ Col X Y B"
  by (meson TS_def assms(1) assms(2) assms(3) col_permutation_5 l9_5 not_col_permutation_1 not_out_bet not_two_sides_id)

lemma l9_18_R2:
  assumes "Col X Y P" and
    "Col A B P" and
    "Bet A P B" and
    "¬ Col X Y A" and
    "¬ Col X Y B"
  shows "X Y TS A B"
  using Col_perm TS_def assms(1) assms(3) assms(4) assms(5) by blast

lemma l9_18:
  assumes "Col X Y P" and
    "Col A B P"
  shows "X Y TS A B  (Bet A P B  ¬ Col X Y A  ¬ Col X Y B)"
  using l9_18_R1 l9_18_R2 assms(1) assms(2) by blast

lemma l9_19_R1:
  assumes "Col X Y P" and
    "Col A B P" and
    "X Y OS A B"
  shows "P Out A B  ¬ Col X Y A"
  by (meson OS_def TS_def assms(1) assms(2) assms(3) col_permutation_5 not_col_permutation_1 not_out_bet one_side_chara)

lemma l9_19_R2:
  assumes "Col X Y P" and
    (*    "Col A B P" and *)
    "P Out A B" and
    "¬ Col X Y A"
  shows "X Y OS A B"
proof -
  obtain D where "X Y TS A D"
    using Col_perm assms(3) l9_10 by blast
  then show ?thesis
    using OS_def assms(1) assms(2) l9_5 not_col_permutation_1 by blast
qed

lemma l9_19:
  assumes "Col X Y P" and
    "Col A B P"
  shows "X Y OS A B  (P Out A B  ¬ Col X Y A)"
  using l9_19_R1 l9_19_R2 assms(1) assms(2) by blast

lemma one_side_not_col123:
  assumes "A B OS X Y"
  shows "¬ Col A B X"
  using assms col_trivial_3 l9_19 by blast

lemma one_side_not_col124:
  assumes "A B OS X Y"
  shows "¬ Col A B Y"
  using assms one_side_not_col123 one_side_symmetry by blast

lemma col_two_sides:
  assumes "Col A B C" and
    "A  C" and
    "A B TS P Q"
  shows "A C TS P Q"
  using assms(1) assms(2) assms(3) col_preserves_two_sides col_trivial_3 by blast

lemma col_one_side:
  assumes "Col A B C" and
    "A  C" and
    "A B OS P Q"
  shows "A C OS P Q"
proof -
  obtain T where "A B TS P T  A B TS Q T" using assms(1) assms(2) assms(3) OS_def by blast
  then show ?thesis
    using col_two_sides OS_def assms(1) assms(2) by blast
qed


lemma out_out_one_side:
  assumes "A B OS X Y" and
    "A Out Y Z"
  shows "A B OS X Z"
  by (meson Col_cases Tarski_neutral_dimensionless.OS_def Tarski_neutral_dimensionless_axioms assms(1) assms(2) col_trivial_3 l9_5)

lemma out_one_side:
  assumes "¬ Col A B X  ¬ Col A B Y" and
    "A Out X Y"
  shows "A B OS X Y"
  using assms(1) assms(2) l6_6 not_col_permutation_2 one_side_reflexivity one_side_symmetry out_out_one_side by blast

lemma bet__ts:
  assumes "A  Y" and
    "¬ Col A B X" and
    "Bet X A Y"
  shows "A B TS X Y"
proof -
  have "¬ Col Y A B"
    using NCol_cases assms(1) assms(2) assms(3) bet_col col2__eq by blast
  then show ?thesis
    by (meson TS_def assms(2) assms(3) col_permutation_3 col_permutation_5 col_trivial_3)
qed

lemma bet_ts__ts:
  assumes "A B TS X Y" and
    "Bet X Y Z"
  shows "A B TS X Z"
proof -
  have "¬ Col Z A B"
    using assms(1) assms(2) bet_col between_equality_2 col_permutation_1 l9_18 by blast
  then show ?thesis
    using TS_def assms(1) assms(2) between_exchange4 by blast
qed

lemma bet_ts__os:
  assumes "A B TS X Y" and
    "Bet X Y Z"
  shows "A B OS Y Z"
  using OS_def assms(1) assms(2) bet_ts__ts l9_2 by blast

lemma l9_31 :
  assumes "A X OS Y Z" and
    "A Z OS Y X"
  shows "A Y TS X Z"
proof -
  have P1: "A  X  A  Z  ¬ Col Y A X  ¬ Col Z A X  ¬ Col Y A Z"
    using assms(1) assms(2) col_permutation_1 one_side_not_col123 one_side_not_col124 os_distincts by blast
  obtain Z' where P2: "Bet Z A Z'  Cong A Z' Z A"
    using segment_construction by blast
  have P3: "Z'  A"
    using P1 P2 cong_diff_4 by blast
  have P4: "A X TS Y Z'"
    by (metis (no_types) P2 P3 assms(1) bet__ts l9_8_2 one_side_not_col124 one_side_symmetry)
  have P5: "¬ Col Y A X"
    using P1 by blast
  obtain T where P6: "Col A T X  Bet Y T Z'"
    using P4 TS_def not_col_permutation_4 by blast
  then have P7: "T  A"
  proof -
    have "¬ Col A Z Y"
      by (simp add: P1 not_col_permutation_1)
    then have f1: "¬ A Out Z Y"
      using out_col by blast
    have "A  Z'"
      using P1 P2 cong_diff_4 by blast
    then show ?thesis
      using f1 by (metis (no_types) P1 P2 P6 l6_2)
  qed
  have P8: "Y A OS Z' T"
    by (smt P1 P2 P3 P6 Tarski_neutral_dimensionless.l6_6 Tarski_neutral_dimensionless_axioms bet_col bet_out col_trivial_2 l6_21 not_col_permutation_1 out_one_side)
  have P9: "A Y TS Z' Z"
    using Col_perm P1 P2 P8 bet__ts between_symmetry one_side_not_col123 by blast
  {
    assume Q0: "Bet T A X"
    have Q1: "Z' Z OS Y T"
      by (metis BetSEq BetS_def P1 P2 P4 P6 TS_def Tarski_neutral_dimensionless.l6_6 Tarski_neutral_dimensionless_axioms bet_col bet_out_1 col_trivial_3 colx not_col_permutation_3 not_col_permutation_4 out_one_side)
    then have Q2: "Z' Out T Y"
      by (metis P6 bet_out_1 os_distincts)
    then have Q3: "A Z OS Y T"
      by (meson Out_cases P1 P2 P6 bet_col col_permutation_3 invert_one_side l9_19_R2)
    have "A Z TS X T"
    proof -
      have R1: "¬ Col X A Z"
        using P1 col_permutation_3 by blast
      have R2: "¬ Col T A Z"
        using Q3 between_trivial one_side_chara by blast
      have " T0. Col T0 A Z  Bet X T0 T"
      proof -
        have S1: "Col A A Z"
          by (simp add: col_trivial_1)
        have "Bet X A T"
          by (simp add: Q0 between_symmetry)
        then show ?thesis using S1 by blast
      qed
      then show ?thesis using R1 R2
        using TS_def by auto
    qed
    have "A Y TS X Z"
      by (meson Q3 Tarski_neutral_dimensionless.l9_8_2 Tarski_neutral_dimensionless.one_side_symmetry Tarski_neutral_dimensionless_axioms A Z TS X T assms(2) l9_9_bis)
  }
  then have P10: "Bet T A X  A Y TS X Z" by blast
  {
    assume R1: "Bet A X T"
    then have R3: "A Y OS Z' X"
      by (meson Bet_cases P1 P6 P8 R1 between_equality invert_one_side not_col_permutation_4 not_out_bet out_out_one_side)
    have "A Y TS X Z"
      using R3 P9 l9_8_2 by blast
  }
  then have P11: "Bet A X T  A Y TS X Z" by blast
  {
    assume R1: "Bet X T A"
    then have R3: "A Y OS T X"
      by (simp add: P5 P7 R1 bet_out_1 not_col_permutation_4 out_one_side)
    then have "A Y TS X Z"
      using P8 P9 invert_two_sides l9_8_2 by blast
  }
  then have "Bet X T A  A Y TS X Z" by blast
  then show ?thesis using P10 P11
    using P6 between_symmetry third_point by blast
qed

lemma col123__nos:
  assumes "Col P Q A"
  shows "¬ P Q OS A B"
  using assms one_side_not_col123 by blast

lemma col124__nos:
  assumes "Col P Q B"
  shows "¬ P Q OS A B"
  using assms one_side_not_col124 by blast

lemma col2_os__os:
  assumes "C  D" and
    "Col A B C" and
    "Col A B D" and
    "A B OS X Y"
  shows "C D OS X Y"
  by (metis assms(1) assms(2) assms(3) assms(4) col3 col_one_side col_trivial_3 invert_one_side os_distincts)

lemma os_out_os:
  assumes "Col A B P" and
    "A B OS C D" and
    "P Out C C'"
  shows "A B OS C' D"
  using OS_def assms(1) assms(2) assms(3) l9_5 not_col_permutation_1 by blast

lemma ts_ts_os:
  assumes "A B TS C D" and
    "C D TS A B"
  shows "A C OS B D"
proof -
  obtain T1 where P1: "Col T1 A B  Bet C T1 D"
    using TS_def assms(1) by blast
  obtain T where P2: "Col T C D  Bet A T B"
    using TS_def assms(2) by blast
  have P3: "T1 = T"
  proof -
    have "A  B"
      using assms(2) ts_distincts by blast
    then show ?thesis
    proof -
      have "Col T1 D C"
        using Col_def P1 by blast
      then have f1: "p. (C = T1  Col C p T1)  ¬ Col C T1 p"
        by (metis assms(1) col_transitivity_1 l6_16_1 ts_distincts)
      have f2: "¬ Col C A B"
        using TS_def assms(1) by presburger
      have f3: "(Bet B T1 A  Bet T1 A B)  Bet A B T1"
        using Col_def P1 by blast
      {
        assume "T1  B"
        then have "C  T1  ¬ Col C T1 B  (p. ¬ Col p T1 B  Col p T1 T)  T  A  T  B"
          using f3 f2 by (metis (no_types) Col_def col_transitivity_1 l6_16_1)
        then have "T  A  T  B  C  T1  ¬ Col C T1 T  T1 = T"
          using f3 by (meson Col_def l6_16_1)
      }
      moreover
      {
        assume "T  A  T  B"
        then have "C  T1  ¬ Col C T1 T  T1 = T"
          using f2 by (metis (no_types) Col_def P1 P2 A  B col_transitivity_1 l6_16_1)
      }
      ultimately have "C  T1  ¬ Col C T1 T  T1 = T"
        using f2 f1 assms(1) ts_distincts by blast
      then show ?thesis
        by (metis (no_types) Col_def P1 P2 assms(1) l6_16_1 ts_distincts)
    qed
  qed
  have P4: "A C OS T B"
    by (metis Col_cases P2 TS_def assms(1) assms(2) bet_out out_one_side)
  then have "C A OS T D"
    by (metis Col_cases P1 TS_def P3 assms(2) bet_out os_distincts out_one_side)
  then show ?thesis
    by (meson P4 Tarski_neutral_dimensionless.invert_one_side Tarski_neutral_dimensionless.one_side_symmetry Tarski_neutral_dimensionless_axioms one_side_transitivity)
qed

lemma col_one_side_out:
  assumes "Col A X Y" and
    "A B OS X Y"
  shows "A Out X Y"
  by (meson assms(1) assms(2) l6_4_2 not_col_distincts not_col_permutation_4 one_side_chara)

lemma col_two_sides_bet:
  assumes "Col A X Y" and
    "A B TS X Y"
  shows "Bet X A Y"
  using Col_cases assms(1) assms(2) l9_8_1 l9_9 or_bet_out out_out_one_side by blast

lemma os_ts1324__os:
  assumes "A X OS Y Z" and
    "A Y TS X Z"
  shows "A Z OS X Y"
proof -
  obtain P where P1: "Col P A Y  Bet X P Z"
    using TS_def assms(2) by blast
  have P2: "A Z OS X P"
    by (metis Col_cases P1 TS_def assms(1) assms(2) bet_col bet_out_1 col124__nos col_trivial_2 l6_6 l9_19)
  have "A Z OS P Y"
  proof -
    have "¬ Col A Z P  ¬ Col A Z Y"
      using P2 col124__nos by blast
    moreover have "A Out P Y"
    proof -
      have "X A OS P Z"
        by (metis Col_cases P1 P2 assms(1) bet_out col123__nos out_one_side)
      then have "A X OS P Y"
        by (meson Tarski_neutral_dimensionless.invert_one_side Tarski_neutral_dimensionless.one_side_symmetry Tarski_neutral_dimensionless_axioms assms(1) one_side_transitivity)
      then show ?thesis
        using P1 col_one_side_out not_col_permutation_4 by blast
    qed
    ultimately show ?thesis
      by (simp add: out_one_side)
  qed
  then show ?thesis
    using P2 one_side_transitivity by blast
qed

lemma ts2__ex_bet2:
  assumes "A C TS B D" and
    "B D TS A C"
  shows " X. Bet A X C  Bet B X D"
  by (metis TS_def assms(1) assms(2) bet_col col_permutation_5 l9_18_R1 not_col_permutation_2)

lemma out_one_side_1:
  assumes "¬ Col A B C" and
    "Col A B X" and
    "X Out C D"
  shows "A B OS C D"
  using assms(1) assms(2) assms(3) not_col_permutation_2 one_side_reflexivity one_side_symmetry os_out_os by blast

lemma out_two_sides_two_sides:
  assumes (*"A ≠ PX" and *)
    "Col A B PX" and
    "PX Out X P" and
    "A B TS P Y"
  shows "A B TS X Y"
  using assms(1) assms(2) assms(3) l6_6 l9_5 not_col_permutation_1 by blast

lemma l8_21_bis:
  assumes "X  Y" and
    "¬ Col C A B"
  shows " P. Cong A P X Y  A B Perp P A  A B TS C P"
proof -
  have P1: "A  B"
    using assms(2) not_col_distincts by blast
  then have " P T. A B Perp P A  Col A B T  Bet C T P"
    using l8_21 by auto
  then obtain P T where P2: "A B Perp P A  Col A B T  Bet C T P" by blast
  have P3: "A B TS C P"
  proof -
    have "¬ Col P A B"
      using P2 col_permutation_1 perp_not_col by blast
    then show ?thesis
      using P2 TS_def assms(2) not_col_permutation_1 by blast
  qed
  have P4: "P  A"
    using P3 ts_distincts by blast
  obtain P' where P5: "(Bet A P P'  Bet A P' P)  Cong A P' X Y"
    using segment_construction_2 P4 by blast
  have P6: "A B Perp P' A"
    by (smt P2 P5 Perp_perm assms(1) bet_col cong_identity cong_symmetry not_bet_distincts not_col_permutation_2 perp_col2)
  have P7: "¬ Col P' A B"
    using NCol_perm P6 col_trivial_3 l8_16_1 by blast
  then have P8: "A B OS P P'"
    by (metis Out_def P4 P5 P6 col_permutation_2 out_one_side perp_not_eq_2)
  then have P9: "A B TS C P'"
    using P3 l9_2 l9_8_2 by blast
  then show ?thesis
    using P5 P6 by blast
qed

lemma ts__ncol:
  assumes "A B TS X Y"
  shows  "¬ Col A X Y  ¬ Col B X Y"
  by (metis TS_def assms col_permutation_1 col_transitivity_2 ts_distincts)

lemma one_or_two_sides_aux:
  assumes "¬ Col C A B" and
    "¬ Col D A B" and
    "Col A C X"
    and "Col B D X"
  shows "A B TS C D  A B OS C D"
proof -
  have P1: "A  X"
    using assms(2) assms(4) col_permutation_2 by blast
  have P2: "B  X"
    using assms(1) assms(3) col_permutation_4 by blast
  have P3: "¬ Col X A B"
    using P1 assms(1) assms(3) col_permutation_5 col_transitivity_1 not_col_permutation_4 by blast
  {
    assume Q0: "Bet A C X  Bet B D X"
    then have Q1: "A B OS C X"
      using assms(1) bet_out not_col_distincts not_col_permutation_1 out_one_side by blast
    then have "A B OS X D"
      by (metis Q0 assms(2) assms(4) bet_out_1 col_permutation_2 col_permutation_3 invert_one_side l6_4_2 not_bet_and_out not_col_distincts out_one_side)
    then have "A B OS C D"
      using Q1 one_side_transitivity by blast
  }
  then have P4: "Bet A C X  Bet B D X  A B OS C D" by blast
  {
    assume "Bet A C X  Bet D X B"
    then have "A B OS C D"
      by (smt P2 assms(1) assms(4) bet_out between_equality_2 l9_10 l9_5 l9_8_1 not_bet_and_out not_col_distincts not_col_permutation_4 out_to_bet out_two_sides_two_sides)
  }
  then have P5: "Bet A C X  Bet D X B  A B OS C D " by blast
  {
    assume Q0: "Bet A C X  Bet X B D"
    have Q1: "A B TS X D"
      using P3 Q0 TS_def assms(2) col_trivial_3 by blast
    have "A B OS X C"
      using Q0 assms(1) bet_out not_col_distincts one_side_reflexivity one_side_symmetry out_out_one_side by blast
    then have "A B TS C D"
      using Q1 l9_8_2 by blast
  }
  then have P6: "Bet A C X  Bet X B D  A B TS C D" by blast
  {
    assume Q1: "Bet C X A  Bet B D X"
    then have Q2: "A B OS C X"
      using P1 assms(1) assms(3) between_equality_2 l6_4_2 not_col_permutation_1 not_col_permutation_4 out_one_side by blast
    have "A B OS X D"
      using Q1 assms(2) bet_out not_col_distincts one_side_reflexivity os_out_os by blast
    then have "A B OS C D" using Q2
      using one_side_transitivity by blast
  }
  then have P7: "Bet C X A  Bet B D X  A B OS C D" by blast
  {
    assume "Bet C X A  Bet D X B"
    then have "A B OS C D"
      by (smt Bet A C X  Bet D X B  A B OS C D Bet C X A  Bet B D X  A B OS C D assms(1) assms(2) assms(3) assms(4) between_symmetry l6_21 l9_18_R2 not_col_distincts ts_ts_os)
  }
  then have P8: "Bet C X A  Bet D X B  A B OS C D" by blast
  {
    assume Q1: "Bet C X A  Bet X B D"
    have Q2: "A B TS X D"
      by (metis P3 Q1 assms(2) bet__ts invert_two_sides not_col_distincts not_col_permutation_3)
    have Q3: "A B OS X C"
      using P1 Q1 assms(1) bet_out_1 not_col_permutation_1 out_one_side by auto
    then have "A B TS C D"
      using Q2 l9_8_2 by blast
  }
  then have P9: "Bet C X A  Bet X B D  A B TS C D" by blast
  {
    assume Q0: "Bet X A C  Bet B D X"
    have Q1: "A B TS X C"
      by (metis P3 Q0 assms(1) bet__ts col_permutation_2 not_col_distincts)
    have "A B OS X D"
      by (metis NCol_cases Q0 Tarski_neutral_dimensionless.out_one_side Tarski_neutral_dimensionless_axioms assms(2) assms(4) bet_out_1 invert_one_side l6_4_1 not_col_distincts not_out_bet)
    then have "A B TS C D"
      using Q1 l9_2 l9_8_2 by blast
  }
  then have P10: "Bet X A C  Bet B D X  A B TS C D" by blast
  {
    assume Q0: "Bet X A C  Bet D X B"
    have Q1: "A B TS X C"
      by (metis NCol_cases P3 Q0 assms(1) bet__ts not_col_distincts)
    have "A B OS X D"
      by (metis P2 P3 Q0 bet_out_1 col_permutation_3 invert_one_side out_one_side)
    then have "A B TS C D"
      using Q1 l9_2 l9_8_2 by blast
  }
  then have P11: "Bet X A C  Bet D X B  A B TS C D"
    by blast
  {
    assume Q0: "Bet X A C  Bet X B D"
    then have Q1: "A B TS C X"
      by (simp add: P1 Q0 assms(1) bet__ts between_symmetry not_col_permutation_1)
    have "A B TS D X"
      by (simp add: P2 Q0 assms(2) bet__ts between_symmetry invert_two_sides not_col_permutation_3)
    then have "A B OS C D"
      using Q1 l9_8_1 by blast
  }
  then have P12: "Bet X A C  Bet X B D  A B OS C D" by blast
  then show ?thesis using P4 P5 P6 P7 P8 P9 P10 P11
    using Col_def assms(3) assms(4) by auto
qed

lemma cop__one_or_two_sides:
  assumes "Coplanar A B C D" and
    "¬ Col C A B" and
    "¬ Col D A B"
  shows "A B TS C D  A B OS C D"
proof -
  obtain X where P1: "Col A B X  Col C D X  Col A C X  Col B D X  Col A D X  Col B C X"
    using Coplanar_def assms(1) by auto
  have P2: "Col A B X  Col C D X  A B TS C D  A B OS C D"
    by (metis TS_def Tarski_neutral_dimensionless.l9_19_R2 Tarski_neutral_dimensionless_axioms assms(2) assms(3) not_col_permutation_3 not_col_permutation_5 not_out_bet)
  have P3: "Col A C X  Col B D X   A B TS C D  A B OS C D"
    using assms(2) assms(3) one_or_two_sides_aux by blast
  have "Col A D X  Col B C X   A B TS C D  A B OS C D"
    using assms(2) assms(3) l9_2 one_or_two_sides_aux one_side_symmetry by blast
  then show ?thesis
    using P1 P2 P3 by blast
qed

lemma os__coplanar:
  assumes "A B OS C D"
  shows "Coplanar A B C D"
proof -
  have P1: "¬ Col A B C"
    using assms one_side_not_col123 by blast
  obtain C' where P2: "Bet C B C'  Cong B C' B C"
    using segment_construction by presburger
  have P3: "A B TS D C'"
    by (metis (no_types) Cong_perm OS_def P2 TS_def assms bet__ts bet_cong_eq invert_one_side l9_10 l9_8_2 one_side_not_col123 ts_distincts)
  obtain T where P4: "Col T A B  Bet D T C'"
    using P3 TS_def by blast
  have P5: "C'  T"
    using P3 P4 TS_def by blast
  have P6: "Col T B C  Coplanar A B C D"
    by (metis Col_def Coplanar_def P2 P4 P5 col_trivial_2 l6_16_1)
  {
    assume Q0: "¬ Col T B C"
    {
      assume R0: "Bet T B A"
      have S1: "B C TS T A"
        by (metis P1 Q0 R0 bet__ts col_permutation_2 not_col_distincts)
      have "C' Out T D"
        using P4 P5 bet_out_1 by auto
      then have "B C OS T D"
        using P2 Q0 bet_col invert_one_side not_col_permutation_3 out_one_side_1 by blast
      then have R1: "B C TS D A"
        using S1 l9_8_2 by blast
      then have "Coplanar A B C D"
        using ncoplanar_perm_9 ts__coplanar by blast
    }
    then have Q1: "Bet T B A  Coplanar A B C D" by blast
    {
      assume R0: "¬ Bet T B A"
      {
        have R2: "B C OS D T"
        proof -
          have S1: "¬ Col B C D"
            by (metis Col_perm P2 P3 P4 Q0 bet_col colx ts_distincts)
          have S2: "Col B C C'"
            by (simp add: P2 bet_col col_permutation_4)
          have S3: "C' Out D T"
            using P4 P5 bet_out_1 l6_6 by auto
          then show ?thesis
            using S1 S2 out_one_side_1 by blast
        qed

        have R3: "B C OS T A"
          using P4 Q0 R0 col_permutation_2 col_permutation_5 not_bet_out out_one_side by blast
      }
      then have R1: "B C OS D A"
        by (metis P2 P4 Q0 bet_col bet_out_1 col_permutation_2 col_permutation_5 os_out_os)
      then have "Coplanar A B C D"
        by (simp add: R1 assms coplanar_perm_19 invert_one_side l9_31 one_side_symmetry ts__coplanar)
    }
    then have "¬ Bet T B A  Coplanar A B C D" by blast
    then have "Coplanar A B C D" using Q1 by blast
  }
  then have "¬ Col T B C  Coplanar A B C D" by blast
  then show ?thesis using P6 by blast
qed

lemma coplanar_trans_1:
  assumes "¬ Col P Q R" and
    "Coplanar P Q R A" and
    "Coplanar P Q R B"
  shows "Coplanar Q R A B"
proof -
  have P1: "Col Q R A  Coplanar Q R A B"
    by (simp add: col__coplanar)
  {
    assume T1: "¬ Col Q R A"
    {
      assume T2: "¬ Col Q R B"
      {
        have "Col Q A B  Coplanar Q R A B"
          using ncop__ncols by blast
        {
          assume S1: "¬ Col Q A B"
          have U1: "Q R TS P A  Q R OS P A"
            by (simp add: T1 assms(1) assms(2) cop__one_or_two_sides coplanar_perm_8 not_col_permutation_2)
          have U2: "Q R TS P B  Q R OS P B"
            using T2 assms(1) assms(3) col_permutation_1 cop__one_or_two_sides coplanar_perm_8 by blast
          have W1: "Q R TS P A  Q R OS P A  Q R TS A B  Q R OS A B"
            using l9_9 by blast
          have W2: "Q R TS P A  Q R OS P B  Q R TS A B  Q R OS A B"
            using l9_2 l9_8_2 by blast
          have W3: "Q R TS P B  Q R OS P A  Q R TS A B  Q R OS A B"
            using l9_8_2 by blast
          have "Q R TS P B  Q R OS P B  Q R TS A B  Q R OS A B"
            using l9_9 by blast
          then have S2: "Q R TS A B  Q R OS A B" using U1 U2 W1 W2 W3
            using OS_def l9_2 one_side_transitivity by blast
          have "Coplanar Q R A B"
            using S2 os__coplanar ts__coplanar by blast
        }
        then have "¬ Col Q A B  Coplanar Q R A B" by blast
      }
      then have "Coplanar Q R A B"
        using ncop__ncols by blast
    }
    then have "¬ Col Q R B  Coplanar Q R A B"
      by blast
  }
  then have "¬ Col Q R A  Coplanar Q R A B"
    using ncop__ncols by blast
  then show ?thesis using P1 by blast
qed

lemma col_cop__cop:
  assumes "Coplanar A B C D" and
    "C  D" and
    "Col C D E"
  shows "Coplanar A B C E"
proof -
  have "Col D A C  Coplanar A B C E"
    by (meson assms(2) assms(3) col_permutation_1 l6_16_1 ncop__ncols)
  moreover
  {
    assume "¬ Col D A C"
    then have "Coplanar A C B E"
      by (meson assms(1) assms(3) col__coplanar coplanar_trans_1 ncoplanar_perm_11 ncoplanar_perm_13)
    then have "Coplanar A B C E"
      using ncoplanar_perm_2 by blast
  }
  ultimately show ?thesis
    by blast
qed

lemma bet_cop__cop:
  assumes "Coplanar A B C E" and
    "Bet C D E"
  shows "Coplanar A B C D"
  by (metis NCol_perm Tarski_neutral_dimensionless.col_cop__cop Tarski_neutral_dimensionless_axioms assms(1) assms(2) bet_col bet_neq12__neq)

lemma col2_cop__cop:
  assumes "Coplanar A B C D" and
    "C  D" and
    "Col C D E" and
    "Col C D F"
  shows "Coplanar A B E F"
proof cases
  assume "C = E"
  then show ?thesis
    using assms(1) assms(2) assms(4) col_cop__cop by blast
next
  assume "C  E"
  then show ?thesis
    by (metis assms(1) assms(2) assms(3) assms(4) col_cop__cop col_transitivity_1 ncoplanar_perm_1 not_col_permutation_4)
qed

lemma col_cop2__cop:
  assumes "U  V" and
    "Coplanar A B C U" and
    "Coplanar A B C V" and
    "Col U V P"
  shows "Coplanar A B C P"
proof cases
  assume "Col A B C"
  then show ?thesis
    using ncop__ncol by blast
next
  assume "¬ Col A B C"
  then show ?thesis
    by (smt Col_perm assms(1) assms(2) assms(3) assms(4) col_cop__cop coplanar_trans_1 ncoplanar_perm_1 ncoplanar_perm_14 ncoplanar_perm_15 ncoplanar_perm_23)
qed

lemma bet_cop2__cop:
  assumes "Coplanar A B C U" and
    "Coplanar A B C W" and
    "Bet U V W"
  shows "Coplanar A B C V"
proof -
  have "Col U V W"
    using assms(3) bet_col by blast
  then have "Col U W V"
    by (meson not_col_permutation_5)
  then show ?thesis
    using assms(1) assms(2) assms(3) bet_neq23__neq col_cop2__cop by blast
qed

lemma coplanar_pseudo_trans:
  assumes "¬ Col P Q R" and
    "Coplanar P Q R A" and
    "Coplanar P Q R B" and
    "Coplanar P Q R C" and
    "Coplanar P Q R D"
  shows "Coplanar A B C D"
proof cases
  have LEM1: "(¬ Col P Q R  Coplanar P Q R A  Coplanar P Q R B  Coplanar P Q R C)  Coplanar A B C R"
    by (smt col_transitivity_2 coplanar_trans_1 ncop__ncols ncoplanar_perm_19 ncoplanar_perm_21)
  assume P2: "Col P Q D"
  have P3: "P  Q"
    using assms(1) col_trivial_1 by blast
  have P4: "Coplanar A B C Q"
    by (smt assms(1) assms(2) assms(3) assms(4) col2_cop__cop coplanar_trans_1 ncoplanar_perm_9 not_col_distincts)
  have P5: "¬ Col Q R P"
    using Col_cases assms(1) by blast
  have P6: "Coplanar Q R P A"
    using assms(2) ncoplanar_perm_12 by blast
  have P7: "Coplanar Q R P B"
    using assms(3) ncoplanar_perm_12 by blast
  have P8: "Coplanar Q R P C"
    using assms(4) ncoplanar_perm_12 by blast
  then have "Coplanar A B C P" using LEM1 P5 P6 P7
    by (smt col_transitivity_2 coplanar_trans_1 ncop__ncols ncoplanar_perm_19)
  then show ?thesis
    using LEM1 P2 P3 P4 col_cop2__cop by blast
next
  assume P9: "¬ Col P Q D"
  have P10: "Coplanar P Q D A"
    using NCol_cases assms(1) assms(2) assms(5) coplanar_trans_1 ncoplanar_perm_8 by blast
  have P11: "Coplanar P Q D B"
    using assms(1) assms(3) assms(5) col_permutation_1 coplanar_perm_12 coplanar_trans_1 by blast
  have "Coplanar P Q D C"
    by (meson assms(1) assms(4) assms(5) coplanar_perm_7 coplanar_trans_1 ncoplanar_perm_14 not_col_permutation_3)
  then show ?thesis using P9 P10 P11
    by (smt P10 P11 P9 col3 coplanar_trans_1 ncop__ncol ncoplanar_perm_20 ncoplanar_perm_23 not_col_distincts)
qed

lemma l9_30:
  assumes "¬ Coplanar A B C P" and
    "¬ Col D E F" and
    "Coplanar D E F P" and
    "Coplanar A B C X" and
    "Coplanar A B C Y" and
    "Coplanar A B C Z" and
    "Coplanar D E F X" and
    "Coplanar D E F Y" and
    "Coplanar D E F Z"
  shows "Col X Y Z"
proof -
  {
    assume P1: "¬ Col X Y Z"
    have P2: "¬ Col A B C"
      using assms(1) col__coplanar by blast
    have "Coplanar A B C P"
    proof -
      have Q2: "Coplanar X Y Z A"
        by (smt P2 assms(4) assms(5) assms(6) col2_cop__cop coplanar_trans_1 ncoplanar_perm_18 not_col_distincts)
      have Q3: "Coplanar X Y Z B"
        using P2 assms(4) assms(5) assms(6) col_trivial_3 coplanar_pseudo_trans ncop__ncols by blast
      have Q4: "Coplanar X Y Z C"
        using P2 assms(4) assms(5) assms(6) col_trivial_2 coplanar_pseudo_trans ncop__ncols by blast
      have "Coplanar X Y Z P"
        using assms(2) assms(3) assms(7) assms(8) assms(9) coplanar_pseudo_trans by blast
      then show ?thesis using P1 Q2 Q3 Q4
        using assms(2) assms(3) assms(7) assms(8) assms(9) coplanar_pseudo_trans by blast
    qed
    then have "False" using assms(1) by blast
  }
  then show ?thesis by blast
qed

lemma cop_per2__col:
  assumes "Coplanar A X Y Z" and
    "A  Z" and
    "Per X Z A" and
    "Per Y Z A"
  shows "Col X Y Z"
proof cases
  assume "X = Y  X = Z  Y = Z"
  then show ?thesis
    using not_col_distincts by blast
next
  assume H1:"¬ (X = Y  X = Z  Y = Z)"
  obtain B where P1: "Cong X A X B  Z Midpoint A B  Cong Y A Y B"
    using Per_def assms(3) assms(4) per_double_cong by blast
  have P2: "X  Y"
    using H1 by blast
  have P3: "X  Z"
    using H1 by blast
  have P4: "Y  Z"
    using H1 by blast
  obtain I where P5: " Col A X I  Col Y Z I 
     Col A Y I  Col X Z I  Col A Z I  Col X Y I"
    using Coplanar_def assms(1) by auto
  have P6: "Col A X I  Col Y Z I  Col X Y Z"
    by (smt P1 P4 assms(2) l4_17 l4_18 l7_13 l7_2 l7_3_2 midpoint_distinct_2 not_col_permutation_1)
  have P7: "Col A Y I  Col X Z I  Col X Y Z"
    by (smt P1 P3 assms(2) col_permutation_1 col_permutation_5 l4_17 l4_18 l7_13 l7_2 l7_3_2 midpoint_distinct_2)
  have "Col A Z I  Col X Y I  Col X Y Z"
    by (metis P1 P2 assms(2) col_permutation_1 l4_17 l4_18 l7_13 l7_2 l7_3_2 midpoint_distinct_2)
  then show ?thesis
    using P5 P6 P7 by blast
qed

lemma cop_perp2__col:
  assumes "Coplanar A B Y Z" and
    "X Y Perp A B" and
    "X Z Perp A B"
  shows "Col X Y Z"
proof cases
  assume P1: "Col A B X"
  {
    assume Q0: "X = A"
    then have Q1: "X  B"
      using assms(3) perp_not_eq_2 by blast
    have Q2: "Coplanar B Y Z X"
      by (simp add: Q0 assms(1) coplanar_perm_9)
    have Q3: "Per Y X B"
      using Q0 assms(2) perp_per_2 by blast
    have "Per Z X B"
      using Q0 assms(3) perp_per_2 by blast
    then have "Col X Y Z"
      using Q1 Q2 Q3 cop_per2__col not_col_permutation_1 by blast
  }
  then have P2: "X = A  Col X Y Z" by blast
  {
    assume Q0: "X  A"
    have Q1: "A X Perp X Y"
      by (metis P1 Perp_perm Q0 assms(2) perp_col1)
    have Q2: "A X Perp X Z"
      by (metis P1 Perp_perm Q0 assms(3) perp_col1)
    have Q3: "Coplanar A Y Z X"
      by (smt P1 assms(1) assms(2) col2_cop__cop col_trivial_3 coplanar_perm_12 coplanar_perm_16 perp_distinct)
    have Q4: "Per Y X A"
      using Perp_perm Q1 perp_per_2 by blast
    have "Per Z X A"
      using P1 Q0 assms(3) perp_col1 perp_per_1 by auto
    then have "Col X Y Z"
      using Q0 Q3 Q4 cop_per2__col not_col_permutation_1 by blast
  }
  then have "X  A  Col X Y Z" by blast
  then show ?thesis
    using P2 by blast
next
  assume P1: "¬ Col A B X"
  obtain Y0 where P2: "Y0 PerpAt X Y A B"
    using Perp_def assms(2) by blast
  obtain Z0 where P3: "Z0 PerpAt X Z A B"
    using Perp_def assms(3) by auto
  have P4: "X Y0 Perp A B"
    by (metis P1 P2 assms(2) perp_col perp_in_col)
  have P5: "X Z0 Perp A B"
    by (metis P1 P3 assms(3) perp_col perp_in_col)
  have P6: "Y0 = Z0"
    by (meson P1 P2 P3 P4 P5 Perp_perm l8_18_uniqueness perp_in_col)
  have P7: "X  Y0"
    using P4 perp_not_eq_1 by blast
  have P8: "Col X Y0 Y"
    using P2 col_permutation_5 perp_in_col by blast
  have "Col X Y0 Z"
    using P3 P6 col_permutation_5 perp_in_col by blast
  then show ?thesis
    using P7 P8 col_transitivity_1 by blast
qed

lemma two_sides_dec:
  shows "A B TS C D  ¬ A B TS C D"
  by simp

lemma cop_nts__os:
  assumes "Coplanar A B C D" and
    "¬ Col C A B" and
    "¬ Col D A B" and
    "¬ A B TS C D"
  shows "A B OS C D"
  using assms(1) assms(2) assms(3) assms(4) cop__one_or_two_sides by blast

lemma cop_nos__ts:
  assumes "Coplanar A B C D" and
    "¬ Col C A B" and
    "¬ Col D A B" and
    "¬ A B OS C D"
  shows "A B TS C D"
  using assms(1) assms(2) assms(3) assms(4) cop_nts__os by blast

lemma one_side_dec:
  "A B OS C D  ¬ A B OS C D"
  by simp

lemma cop_dec:
  "Coplanar A B C D  ¬ Coplanar A B C D"
  by simp

lemma ex_diff_cop:
  " E. Coplanar A B C E  D  E"
  by (metis col_trivial_2 diff_col_ex ncop__ncols)

lemma ex_ncol_cop:
  assumes "D  E"
  shows " F. Coplanar A B C F  ¬ Col D E F"
proof cases
  assume "Col A B C"
  then show ?thesis
    using assms ncop__ncols not_col_exists by blast
next
  assume P1: "¬ Col A B C"
  then show ?thesis
  proof -
    have P2: "(Col D E A  Col D E B)  ( F. Coplanar A B C F  ¬ Col D E F)"
      by (meson P1 assms col3 col_trivial_2 ncop__ncols)
    have P3: "(¬Col D E A  Col D E B)  ( F. Coplanar A B C F  ¬ Col D E F)"
      using col_trivial_3 ncop__ncols by blast
    have P4: "(Col D E A  ¬Col D E B)  ( F. Coplanar A B C F  ¬ Col D E F)"
      using col_trivial_2 ncop__ncols by blast
    have "(¬Col D E A  ¬Col D E B)  ( F. Coplanar A B C F  ¬ Col D E F)"
      using col_trivial_3 ncop__ncols by blast
    then show ?thesis using P2 P3 P4 by blast
  qed
qed

lemma ex_ncol_cop2:
  " E F. (Coplanar A B C E  Coplanar A B C F  ¬ Col D E F)"
proof -
  have f1: "p pa pb. Coplanar pb pa p pb"
    by (meson col_trivial_3 ncop__ncols)
  have f2: "p pa pb. Coplanar pb pa p p"
    by (meson Col_perm col_trivial_3 ncop__ncols)
  obtain pp :: "'p  'p  'p" where
    f3: "p pa. p = pa  ¬ Col p pa (pp p pa)"
    using not_col_exists by moura
  have f4: "p pa pb. Coplanar pb pb pa p"
    by (meson Col_perm col_trivial_3 ncop__ncols)
  have "p. A  p"
    by (meson col_trivial_3 diff_col_ex3)
  moreover
  { assume "B  A"
    then have "D = B  (p. ¬ Col D p A  Coplanar A B C p)"
      using f3 f2 by (metis (no_types) Col_perm ncop__ncols)
    then have "D = B  (p pa. Coplanar A B C p  Coplanar A B C pa  ¬ Col D p pa)"
      using f1 by blast }
  moreover
  { assume "D  B"
    moreover
    { assume "p. D  B  ¬ Coplanar A B C p"
      then have "D  B  ¬ Col A B C"
        using ncop__ncols by blast
      then have "p. ¬ Col D p B  Coplanar A B C p"
        using f2 f1 by (metis (no_types) Col_perm col_transitivity_1) }
    ultimately have ?thesis
      using f3 by (metis (no_types) col_trivial_3 ncop__ncols) }
  ultimately show ?thesis
    using f4 f3 by blast
qed

lemma col2_cop2__eq:
  assumes "¬ Coplanar A B C U" and
    "U  V" and
    "Coplanar A B C P" and
    "Coplanar A B C Q" and
    "Col U V P" and
    "Col U V Q"
  shows "P = Q"
proof -
  have "Col U Q P"
    by (meson assms(2) assms(5) assms(6) col_transitivity_1)
  then have "Col P Q U"
    using not_col_permutation_3 by blast
  then show ?thesis
    using assms(1) assms(3) assms(4) col_cop2__cop by blast
qed

lemma cong3_cop2__col:
  assumes "Coplanar A B C P" and
    "Coplanar A B C Q" and
    "P  Q" and
    "Cong A P A Q" and
    "Cong B P B Q" and
    "Cong C P C Q"
  shows "Col A B C"
proof cases
  assume "Col A B C"
  then show ?thesis by blast
next
  assume P1: "¬ Col A B C"
  obtain M where P2: "M Midpoint P Q"
    using assms(6) l7_25 by blast
  have P3: "Per A M P"
    using P2 Per_def assms(4) by blast
  have P4: "Per B M P"
    using P2 Per_def assms(5) by blast
  have P5: "Per C M P"
    using P2 Per_def assms(6) by blast
  have "False"
  proof cases
    assume Q1: "A = M"
    have Q2: "Coplanar P B C A"
      using assms(1) ncoplanar_perm_21 by blast
    have Q3: "P  A"
      by (metis assms(3) assms(4) cong_diff_4)
    have Q4: "Per B A P"
      by (simp add: P4 Q1)
    have Q5: "Per C A P"
      by (simp add: P5 Q1)
    then show ?thesis using Q1 Q2 Q3 Q4 cop_per2__col
      using P1 not_col_permutation_1 by blast
  next
    assume Q0: "A  M"
    have Q1: "Col A B M"
    proof -
      have R1: "Coplanar A B P Q"
        using P1 assms(1) assms(2) coplanar_trans_1 ncoplanar_perm_8 not_col_permutation_2 by blast
      then have R2: "Coplanar P A B M"
        using P2 bet_cop__cop coplanar_perm_14 midpoint_bet ncoplanar_perm_6 by blast
      have R3: "P  M"
        using P2 assms(3) l7_3_2 l7_9_bis by blast
      have R4: "Per A M P"
        by (simp add: P3)
      have R5: "Per B M P"
        by (simp add: P4)
      then show ?thesis
        using R2 R3 R4 cop_per2__col by blast
    qed
    have "Col A C M"
    proof -
      have R1: "Coplanar P A C M"
        using P1 Q1 assms(1) col2_cop__cop coplanar_perm_22 ncoplanar_perm_3 not_col_distincts by blast
      have R2: "P  M"
        using P2 assms(3) l7_3_2 symmetric_point_uniqueness by blast
      have R3: "Per A M P"
        by (simp add: P3)
      have "Per C M P"
        by (simp add: P5)
      then show ?thesis
        using R1 R2 R3 cop_per2__col by blast
    qed
    then show ?thesis
      using NCol_perm P1 Q0 Q1 col_trivial_3 colx by blast
  qed
  then show ?thesis by blast
qed

lemma l9_38:
  assumes "A B C TSP P Q"
  shows "A B C TSP Q P"
  using Bet_perm TSP_def assms by blast

lemma l9_39:
  assumes "A B C TSP P R" and
    "Coplanar A B C D" and
    "D Out P Q"
  shows "A B C TSP Q R"
proof -
  have P1: "¬ Col A B C"
    using TSP_def assms(1) ncop__ncol by blast
  have P2: "¬ Coplanar A B C Q"
    by (metis TSP_def assms(1) assms(2) assms(3) col_cop2__cop l6_6 out_col out_diff2)
  have P3: "¬ Coplanar A B C R"
    using TSP_def assms(1) by blast
  obtain T where P3A: "Coplanar A B C T  Bet P T R"
    using TSP_def assms(1) by blast
  have W1: "D = T  A B C TSP Q R"
    using P2 P3 P3A TSP_def assms(3) bet_out__bet by blast
  {
    assume V1: "D  T"
    have V1A: "¬ Col P D T" using P3A col_cop2__cop
      by (metis TSP_def V1 assms(1) assms(2) col2_cop2__eq col_trivial_2)
    have V1B: "D T TS P R"
      by (metis P3 P3A V1A bet__ts invert_two_sides not_col_permutation_3)
    have "D T OS P Q"
      using V1A assms(3) not_col_permutation_1 out_one_side by blast
    then have V2: "D T TS Q R"
      using V1B l9_8_2 by blast
    then obtain T' where V3: "Col T' D T  Bet Q T' R"
      using TS_def by blast
    have V4: "Coplanar A B C T'"
      using Col_cases P3A V1 V3 assms(2) col_cop2__cop by blast
    then have "A B C TSP Q R"
      using P2 P3 TSP_def V3 by blast
  }
  then have "D  T  A B C TSP Q R" by blast
  then show ?thesis using W1 by blast
qed

lemma l9_41_1:
  assumes "A B C TSP P R" and
    "A B C TSP Q R"
  shows "A B C OSP P Q"
  using OSP_def assms(1) assms(2) by blast

lemma l9_41_2:
  assumes "A B C TSP P R" and
    "A B C OSP P Q"
  shows "A B C TSP Q R"
proof -
  have P1: "¬ Coplanar A B C P"
    using TSP_def assms(1) by blast
  obtain S where P2: " A B C TSP P S  A B C TSP Q S"
    using OSP_def assms(2) by blast
  obtain X where P3: "Coplanar A B C X  Bet P X S"
    using P2 TSP_def by blast
  have P4: "¬ Coplanar A B C P  ¬ Coplanar A B C S"
    using P2 TSP_def by blast
  obtain Y where P5: "Coplanar A B C Y  Bet Q Y S"
    using P2 TSP_def by blast
  have P6: "¬ Coplanar A B C Q  ¬ Coplanar A B C S"
    using P2 TSP_def by blast
  have P7: "X  P  S  X  Q  Y  S  Y"
    using P3 P4 P5 P6 by blast
  {
    assume Q1: "Col P Q S"
    have Q2: "X = Y"
    proof -
      have R2: "Q  S"
        using P5 P6 bet_neq12__neq by blast
      have R5: "Col Q S X"
        by (smt Col_def P3 Q1 between_inner_transitivity between_symmetry col_transitivity_2)
      have "Col Q S Y"
        by (simp add: P5 bet_col col_permutation_5)
      then show ?thesis
        using P2 P3 P5 R2 R5 TSP_def col2_cop2__eq by blast
    qed
    then have "X Out P Q"
      by (metis P3 P5 P7 l6_2)
    then have "A B C TSP Q R"
      using P3 assms(1) l9_39 by blast
  }
  then have P7: "Col P Q S  A B C TSP Q R" by blast
  {
    assume Q1: "¬ Col P Q S"
    obtain Z where Q2: "Bet X Z Q  Bet Y Z P"
      using P3 P5 inner_pasch by blast
    {
      assume "X = Z"
      then have "False"
        by (metis P2 P3 P5 Q1 Q2 TSP_def bet_col col_cop2__cop l6_16_1 not_col_permutation_5)
    }
    then have Q3: "X  Z" by blast
    have "Y  Z"
    proof -
      have "X  Z"
        by (meson X = Z  False)
      then have "Z  Y"
        by (metis (no_types) P2 P3 P5 Q2 TSP_def bet_col col_cop2__cop)
      then show ?thesis
        by meson
    qed
    then have "Y Out P Z"
      using Q2 bet_out l6_6 by auto
    then have Q4: "A B C TSP Z R"
      using assms(1) P5 l9_39 by blast
    have "X Out Z Q"
      using Q2 Q3 bet_out by auto
    then have "A B C TSP Q R"
      using Q4 P3 l9_39 by blast
  }
  then have "¬ Col P Q S  A B C TSP Q R" by blast
  then show ?thesis using P7 by blast
qed

lemma tsp_exists:
  assumes "¬ Coplanar A B C P"
  shows " Q. A B C TSP P Q"
proof -
  obtain Q where P1: "Bet P A Q  Cong A Q A P"
    using segment_construction by blast
  have P2: "Coplanar A B C A"
    using coplanar_trivial ncoplanar_perm_5 by blast
  have P3: "¬ Coplanar A B C P"
    by (simp add: assms)
  have P4: "¬ Coplanar A B C Q"
    by (metis P1 P2 Tarski_neutral_dimensionless.col_cop2__cop Tarski_neutral_dimensionless_axioms assms bet_col cong_diff_4 not_col_permutation_2)
  then show ?thesis
    using P1 P2 TSP_def assms by blast
qed


lemma osp_reflexivity:
  assumes "¬ Coplanar A B C P"
  shows "A B C OSP P P"
  by (meson assms l9_41_1 tsp_exists)

lemma osp_symmetry:
  assumes "A B C OSP P Q"
  shows "A B C OSP Q P"
  using OSP_def assms by auto

lemma osp_transitivity:
  assumes "A B C OSP P Q" and
    "A B C OSP Q R"
  shows "A B C OSP P R"
  using OSP_def assms(1) assms(2) l9_41_2 by blast

lemma cop3_tsp__tsp:
  assumes "¬ Col D E F" and
    "Coplanar A B C D" and
    "Coplanar A B C E" and
    "Coplanar A B C F" and
    "A B C TSP P Q"
  shows "D E F TSP P Q"
proof -
  obtain T where P1: "Coplanar A B C T  Bet P T Q"
    using TSP_def assms(5) by blast
  have P2: "¬ Col A B C"
    using TSP_def assms(5) ncop__ncols by blast
  have P3: "Coplanar D E F A  Coplanar D E F B  Coplanar D E F C  Coplanar D E F T"
  proof -
    have P3A: "Coplanar D E F A"
      using P2 assms(2) assms(3) assms(4) col_trivial_3 coplanar_pseudo_trans ncop__ncols by blast
    have P3B: "Coplanar D E F B"
      using P2 assms(2) assms(3) assms(4) col_trivial_2 coplanar_pseudo_trans ncop__ncols by blast
    have P3C: "Coplanar D E F C"
      by (meson P2 assms(2) assms(3) assms(4) coplanar_perm_16 coplanar_pseudo_trans coplanar_trivial)
    have "Coplanar D E F T"
      using P1 P2 assms(2) assms(3) assms(4) coplanar_pseudo_trans by blast
    then show ?thesis using P3A P3B P3C by simp
  qed
  have P4: "¬ Coplanar D E F P"
    using P3 TSP_def assms(1) assms(5) coplanar_pseudo_trans by auto
  have P5: "¬ Coplanar D E F Q"
    by (metis P1 P3 P4 TSP_def assms(5) bet_col bet_col1 col2_cop2__eq)
  have P6: "Coplanar D E F T"
    by (simp add: P3)
  have "Bet P T Q"
    by (simp add: P1)
  then show ?thesis
    using P4 P5 P6 TSP_def by blast
qed

lemma cop3_osp__osp:
  assumes "¬ Col D E F" and
    "Coplanar A B C D" and
    "Coplanar A B C E" and
    "Coplanar A B C F" and
    "A B C OSP P Q"
  shows "D E F OSP P Q"
proof -
  obtain R where P1: "A B C TSP P R  A B C TSP Q R"
    using OSP_def assms(5) by blast
  then show ?thesis
    using OSP_def assms(1) assms(2) assms(3) assms(4) cop3_tsp__tsp by blast
qed

lemma ncop_distincts:
  assumes "¬ Coplanar A B C D"
  shows "A  B  A  C  A  D  B  C  B  D  C  D"
  using Coplanar_def assms col_trivial_1 col_trivial_2 by blast

lemma tsp_distincts:
  assumes "A B C TSP P Q"
  shows "A  B  A  C  B  C  A  P  B  P  C  P  A  Q  B  Q  C  Q  P  Q"
proof -
  obtain pp :: "'p  'p  'p  'p  'p  'p" where
    "x0 x1 x2 x3 x4. (v5. Coplanar x4 x3 x2 v5  Bet x1 v5 x0) = (Coplanar x4 x3 x2 (pp x0 x1 x2 x3 x4)  Bet x1 (pp x0 x1 x2 x3 x4) x0)"
    by moura
  then have f1: "¬ Coplanar A B C P  ¬ Coplanar A B C Q  Coplanar A B C (pp Q P C B A)  Bet P (pp Q P C B A) Q"
    using TSP_def assms by presburger
  then have "Q  pp Q P C B A"
    by force
  then show ?thesis
    using f1 by (meson bet_neq32__neq ncop_distincts)
qed

lemma osp_distincts:
  assumes "A B C OSP P Q"
  shows "A  B  A  C  B  C  A  P  B  P  C  P  A  Q  B  Q  C  Q"
  using OSP_def assms tsp_distincts by blast

lemma tsp__ncop1:
  assumes "A B C TSP P Q"
  shows "¬ Coplanar A B C P"
  using TSP_def assms by blast

lemma tsp__ncop2:
  assumes "A B C TSP P Q"
  shows "¬ Coplanar A B C Q"
  using TSP_def assms by auto

lemma osp__ncop1:
  assumes "A B C OSP P Q"
  shows "¬ Coplanar A B C P"
  using OSP_def TSP_def assms by blast

lemma osp__ncop2:
  assumes "A B C OSP P Q"
  shows "¬ Coplanar A B C Q"
  using assms osp__ncop1 osp_symmetry by blast

lemma tsp__nosp:
  assumes "A B C TSP P Q"
  shows "¬ A B C OSP P Q"
  using assms l9_41_2 tsp_distincts by blast

lemma osp__ntsp:
  assumes "A B C OSP P Q"
  shows "¬ A B C TSP P Q"
  using assms tsp__nosp by blast

lemma osp_bet__osp:
  assumes "A B C OSP P R" and
    "Bet P Q R"
  shows "A B C OSP P Q"
proof -
  obtain S where P1: "A B C TSP P S"
    using OSP_def assms(1) by blast
  then obtain Y where P2: "Coplanar A B C Y  Bet R Y S"
    using TSP_def assms(1) l9_41_2 by blast
  obtain X where Q1: "Coplanar A B C X  Bet P X S"
    using P1 TSP_def by blast
  have Q2: "P  X  S  X  R  Y"
    using P1 P2 Q1 TSP_def assms(1) osp__ncop2 by auto
  {
    assume P3: "Col P R S"
    have P5: "A B C TSP Q S"
    proof -
      have Q3: "X = Y"
      proof -
        have R1: "¬ Coplanar A B C R"
          using assms(1) osp__ncop2 by blast
        have R2: "R  S"
          using P1 assms(1) osp__ntsp by blast
        have R5: "Col R S X"
          by (smt Col_def P3 Q1 bet_col1 between_exchange4 between_symmetry)
        have "Col R S Y"
          using P2 bet_col col_permutation_5 by blast
        then show ?thesis
          using R1 R2 Q1 P2 R5 col2_cop2__eq by blast
      qed
      then have "Y Out P Q"
        by (smt P2 P3 Q1 Q2 assms(2) bet_col1 between_exchange4 between_symmetry l6_3_2 l6_4_2 not_bet_and_out third_point)
      then show ?thesis
        using P1 P2 l9_39 by blast
    qed
    then have "A B C OSP P Q"
      using OSP_def P1 P2 l9_39 by blast
  }
  then have H1: "Col P R S  A B C OSP P Q" by blast
  {
    assume T1: "¬ Col P R S"
    have T2: "X Y OS P R"
    proof -
      have T3: "P  X  S  X  R  Y  S  Y"
        using P1 P2 Q2 TSP_def by auto
      have T4: "¬ Col S X Y"
        by (metis P2 Q1 T1 T3 bet_out_1 col_out2_col col_permutation_5 not_col_permutation_4)
      have T5: "X Y TS P S"
        by (metis Col_perm Q1 Q2 T4 bet__ts bet_col col_transitivity_2)
      have T6: "X Y TS R S"
        by (metis P2 Q1 T4 assms(1) bet__ts col_cop2__cop invert_two_sides not_col_distincts osp__ncop2)
      then show ?thesis
        using T5 l9_8_1 by auto
    qed
    then have T7: "X Y OS P Q"
      using assms(2) l9_17 by blast
    then obtain S' where T7A: "X Y TS P S'  X Y TS Q S'"
      using OS_def by blast
    have T7B: "¬ Col P X Y  ¬ Col S' X Y  ( T::'p. Col T X Y  Bet P T S')"
      using T7A TS_def by auto
    have T7C: "¬ Col Q X Y  ¬ Col S' X Y  ( T::'p. Col T X Y  Bet Q T S')"
      using T7A TS_def by blast
    obtain X' where T9: "Col X' X Y  Bet P X' S'  X Y TS Q S'"
      using T7A T7B by blast
    obtain Y' where T10: "Col Y' X Y  Bet Q Y' S'"
      using T7C by blast
    have T11: "Coplanar A B C X'"
      using Col_cases P2 Q1 T9 col_cop2__cop ts_distincts by blast
    have T12: "Coplanar A B C Y'"
      using Col_cases P2 Q1 T10 T9 col_cop2__cop ts_distincts by blast
    have T13: "¬ Coplanar A B C S'"
      using T11 T7C T9 assms(1) bet_col bet_col1 col2_cop2__eq osp__ncop1 by fastforce
    then have "A B C OSP P Q"
    proof -
      have R1: "A B C TSP P S'"
        using P1 T11 T13 T9 TSP_def by blast
      have "A B C TSP Q S'"
        by (metis T10 T12 T13 T7C TSP_def bet_col col_cop2__cop)
      then show ?thesis using R1 by (smt l9_41_1)
    qed
  }
  then show ?thesis using H1 by blast
qed

lemma l9_18_3:
  assumes "Coplanar A B C P" and
    "Col X Y P"
  shows "A B C TSP X Y  (Bet X P Y  ¬ Coplanar A B C X  ¬ Coplanar A B C Y)"
  by (meson TSP_def assms(1) assms(2) l9_39 not_bet_out not_col_permutation_5 tsp_distincts)

lemma bet_cop__tsp:
  assumes "¬ Coplanar A B C X" and
    "P  Y" and
    "Coplanar A B C P" and
    "Bet X P Y"
  shows "A B C TSP X Y"
  using TSP_def assms(1) assms(2) assms(3) assms(4) bet_col bet_col1 col2_cop2__eq by fastforce

lemma cop_out__osp:
  assumes "¬ Coplanar A B C X" and
    "Coplanar A B C P" and
    "P Out X Y"
  shows "A B C OSP X Y"
  by (meson OSP_def assms(1) assms(2) assms(3) l9_39 tsp_exists)

lemma l9_19_3:
  assumes "Coplanar A B C P" and
    "Col X Y P"
  shows "A B C OSP X Y  (P Out X Y  ¬ Coplanar A B C X)"
  by (meson assms(1) assms(2) cop_out__osp l6_4_2 l9_18_3 not_col_permutation_5 osp__ncop1 osp__ncop2 tsp__nosp)

lemma cop2_ts__tsp:
  assumes "¬ Coplanar A B C X" and "Coplanar A B C D" and
    "Coplanar A B C E" and "D E TS X Y"
  shows "A B C TSP X Y"
proof -
  obtain T where P1: "Col T D E  Bet X T Y"
    using TS_def assms(4) by blast
  have P2: "Coplanar A B C T"
    using P1 assms(2) assms(3) assms(4) col_cop2__cop not_col_permutation_2 ts_distincts by blast
  then show ?thesis
    by (metis P1 TS_def assms(1) assms(4) bet_cop__tsp)
qed

lemma cop2_os__osp:
  assumes "¬ Coplanar A B C X" and
    "Coplanar A B C D" and
    "Coplanar A B C E" and
    "D E OS X Y"
  shows "A B C OSP X Y"
proof -
  obtain Z where P1: "D E TS X Z  D E TS Y Z"
    using OS_def assms(4) by blast
  then have P2: "A B C TSP X Z"
    using assms(1) assms(2) assms(3) cop2_ts__tsp by blast
  then have P3: "A B C TSP Y Z"
    by (meson P1 assms(2) assms(3) cop2_ts__tsp l9_2 tsp__ncop2)
  then show ?thesis
    using P2 l9_41_1 by blast
qed

lemma cop3_tsp__ts:
  assumes "D  E" and
    "Coplanar A B C D" and
    "Coplanar A B C E" and
    "Coplanar D E X Y" and
    "A B C TSP X Y"
  shows "D E TS X Y"
  by (meson assms(1) assms(2) assms(3) assms(4) assms(5) col_cop2__cop cop2_os__osp cop_nts__os not_col_permutation_2 tsp__ncop1 tsp__ncop2 tsp__nosp)

lemma cop3_osp__os:
  assumes "D  E" and
    "Coplanar A B C D" and
    "Coplanar A B C E" and
    "Coplanar D E X Y" and
    "A B C OSP X Y"
  shows "D E OS X Y"
  by (meson assms(1) assms(2) assms(3) assms(4) assms(5) col_cop2__cop cop2_ts__tsp cop_nts__os not_col_permutation_2 osp__ncop1 osp__ncop2 tsp__nosp)

lemma cop_tsp__ex_cop2:
  assumes (*"Coplanar A B C P" and*)
    "A B C TSP D E"
  shows " Q. (Coplanar A B C Q  Coplanar D E P Q  P  Q)"
proof cases
  assume "Col D E P"
  then show ?thesis
    by (meson ex_diff_cop ncop__ncols)
next
  assume "¬ Col D E P"
  then obtain Q where "Coplanar A B C Q  Bet D Q E  ¬ Col D E P"
    using TSP_def assms(1) by blast
  then show ?thesis
    using Col_perm bet_col ncop__ncols by blast
qed

lemma cop_osp__ex_cop2:
  assumes "Coplanar A B C P" and
    "A B C OSP D E"
  shows " Q. Coplanar A B C Q  Coplanar D E P Q  P  Q"
proof cases
  assume "Col D E P"
  then show ?thesis
    by (metis col_trivial_3 diff_col_ex ncop__ncols)
next
  assume P1: "¬ Col D E P"
  obtain E' where P2: "Bet E P E'  Cong P E' P E"
    using segment_construction by blast
  have P3: "¬ Col D E' P"
    by (metis P1 P2 bet_col bet_cong_eq between_symmetry col_permutation_5 l5_2 l6_16_1)
  have P4: "A B C TSP D E'"
    by (metis P2 P3 assms(1) assms(2) bet_cop__tsp l9_41_2 not_col_distincts osp__ncop2 osp_symmetry)
  then have "¬ Coplanar A B C D  ¬ Coplanar A B C E'  ( T. Coplanar A B C T  Bet D T E')"
    by (simp add: TSP_def)
  then obtain Q where P7: "Coplanar A B C Q  Bet D Q E'"
    by blast
  then have "Coplanar D E' P Q"
    using bet_col ncop__ncols ncoplanar_perm_5 by blast
  then have "Coplanar D E P Q"
    using Col_perm P2 P3 bet_col col_cop__cop ncoplanar_perm_5 not_col_distincts by blast
  then show ?thesis
    using P3 P7 bet_col col_permutation_5 by blast
qed

lemma sac__coplanar:
  assumes "Saccheri A B C D"
  shows "Coplanar A B C D"
  using Saccheri_def assms ncoplanar_perm_4 os__coplanar by blast

subsection "Line reflexivity"

subsubsection "Dimensionless"

lemma Ch10_Goal1:
  assumes "¬ Coplanar D C B A"
  shows "¬ Coplanar A B C D"
  by (simp add: assms ncoplanar_perm_23)

lemma ex_sym:
  " Y. (A B Perp X Y  X = Y)  ( M. Col A B M  M Midpoint X Y)"
proof cases
  assume "Col A B X"
  thus ?thesis
    using l7_3_2 by blast
next
  assume "¬ Col A B X"
  then obtain M0 where P1: "Col A B M0  A B Perp X M0"
    using l8_18_existence by blast
  obtain Z where P2: "M0 Midpoint X Z"
    using symmetric_point_construction by blast
  thus ?thesis
    by (metis (full_types) P1 Perp_cases bet_col midpoint_bet perp_col)
qed

lemma is_image_is_image_spec:
  assumes "A  B"
  shows "P' P Reflect A B  P' P ReflectL A B"
  by (simp add: Reflect_def assms)

lemma ex_sym1:
  assumes "A  B"
  shows " Y. (A B Perp X Y  X = Y)  ( M. Col A B M  M Midpoint X Y  X Y Reflect A B)"
proof cases
  assume "Col A B X"
  thus ?thesis
    by (meson ReflectL_def Reflect_def assms l7_3_2)
next
  assume P0: "¬ Col A B X"
  then obtain M0 where P1: "Col A B M0  A B Perp X M0"
    using l8_18_existence by blast
  obtain Z where P2: "M0 Midpoint X Z"
    using symmetric_point_construction by blast
  have P3: "A B Perp X Z"
  proof cases
    assume "X = Z"
    thus ?thesis
      using P1 P2 P0 midpoint_distinct by blast
  next
    assume "X  Z"
    then have P2: "X Z Perp A B"
      using P1 P2 Perp_cases bet_col midpoint_bet perp_col by blast
    show ?thesis
      by (simp add: Tarski_neutral_dimensionless.Perp_perm Tarski_neutral_dimensionless_axioms P2)
  qed
  have P10: "(A B Perp X Z  X = Z)"
    by (simp add: P3)
  have " M. Col A B M  M Midpoint X Z  X Z Reflect A B"
    using P1 P2 P3 ReflectL_def assms is_image_is_image_spec l7_2 perp_right_comm by blast
  thus ?thesis
    using P3 by blast
qed

lemma l10_2_uniqueness:
  assumes "P1 P Reflect A B" and
    "P2 P Reflect A B"
  shows "P1 = P2"
proof cases
  assume "A = B"
  thus ?thesis
    using Reflect_def assms(1) assms(2) symmetric_point_uniqueness by auto
next
  assume P1: "A  B"
  have P1A: "P1 P ReflectL A B"
    using P1 assms(1) is_image_is_image_spec by auto
  then have P1B: "A B Perp P P1  P = P1"
    using ReflectL_def by blast
  have P2A: "P2 P ReflectL A B"
    using P1 assms(2) is_image_is_image_spec by auto
  then have P2B: "A B Perp P P2  P = P2"
    using ReflectL_def by blast
  obtain X where R1: "X Midpoint P P1  Col A B X"
    by (metis ReflectL_def assms(1) col_trivial_1 is_image_is_image_spec midpoint_existence)
  obtain Y where R2: "Y Midpoint P P2  Col A B Y"
    by (metis ReflectL_def assms(2) col_trivial_1 is_image_is_image_spec midpoint_existence)
  {
    assume Q1:"(A B Perp P P1  A B Perp P P2)"
    have S1: "P  X"
    proof -
      {
        assume "P = X"
        then have "P = P1"
          using R1 is_midpoint_id by blast
        then have "A B Perp P P"
          using Q1 by blast
        then have "False"
          using perp_distinct by blast
      }
      thus ?thesis by blast
    qed
    then have "P1 = P2"
      by (smt Perp_cases Q1 thesis. (X. X Midpoint P P1  Col A B X  thesis)  thesis thesis. (Y. Y Midpoint P P2  Col A B Y  thesis)  thesis col_permutation_1 l7_2 l7_9 l8_18_uniqueness midpoint_col perp_col perp_not_col2)
  }
  then have T1: "(A B Perp P P1  A B Perp P P2)  P1 = P2" by blast
  have T2: "(P = P1  A B Perp P P2)  P1 = P2"
    by (metis R1 R2 col3 col_trivial_2 col_trivial_3 midpoint_col midpoint_distinct_1 midpoint_distinct_2 perp_not_col2)
  have T3: "(P = P2  A B Perp P P1)  P1 = P2"
    by (metis R1 R2 col_trivial_2 midpoint_col midpoint_distinct_3 perp_col2 perp_not_col2)
  thus ?thesis
    using T1 T2 T3 P1B P2B by blast
qed

lemma l10_2_uniqueness_spec:
  assumes "P1 P ReflectL A B" and
    "P2 P ReflectL A B"
  shows "P1 = P2"
proof -
  have "A B Perp P P1  P = P1"
    using ReflectL_def assms(1) by blast
  moreover obtain X1 where "X1 Midpoint P P1  Col A B X1"
    using ReflectL_def assms(1) by blast
  moreover have "A B Perp P P2  P = P2"
    using ReflectL_def assms(2) by blast
  moreover obtain X2 where "X2 Midpoint P P2  Col A B X2"
    using ReflectL_def assms(2) by blast
  ultimately show ?thesis
    by (smt col_permutation_1 l8_16_1 l8_18_uniqueness midpoint_col midpoint_distinct_3 perp_col1 symmetric_point_uniqueness)
qed

lemma l10_2_existence_spec:
  " P'. P' P ReflectL A B"
proof cases
  assume "Col A B P"
  thus ?thesis
    using ReflectL_def l7_3_2 by blast
next
  assume "¬ Col A B P"
  then obtain X where "Col A B X  A B Perp P X"
    using l8_18_existence by blast
  moreover obtain P' where "X Midpoint P P'"
    using symmetric_point_construction by blast
  ultimately show ?thesis
    using ReflectL_def bet_col midpoint_bet perp_col1 by blast
qed

lemma l10_2_existence:
  " P'. P' P Reflect A B"
  by (metis Reflect_def l10_2_existence_spec symmetric_point_construction)

lemma l10_4_spec:
  assumes "P P' ReflectL A B"
  shows "P' P ReflectL A B"
proof -
  obtain X where "X Midpoint P P'  Col A B X"
    using ReflectL_def assms l7_2 by blast
  thus ?thesis
    using Perp_cases ReflectL_def assms by auto
qed

lemma l10_4:
  assumes "P P' Reflect A B"
  shows "P' P Reflect A B"
  using Reflect_def Tarski_neutral_dimensionless.l7_2 Tarski_neutral_dimensionless_axioms assms l10_4_spec by fastforce

lemma l10_5:
  assumes "P' P Reflect A B" and
    "P'' P' Reflect A B"
  shows "P = P''"
  by (meson assms(1) assms(2) l10_2_uniqueness l10_4)

lemma l10_6_uniqueness:
  assumes "P P1 Reflect A B" and
    "P P2 Reflect A B"
  shows "P1 = P2"
  using assms(1) assms(2) l10_4 l10_5 by blast

lemma l10_6_uniqueness_spec:
  assumes "P P1 ReflectL A B" and
    "P P2 ReflectL A B"
  shows "P1 = P2"
  using assms(1) assms(2) l10_2_uniqueness_spec l10_4_spec by blast

lemma l10_6_existence_spec:
  assumes "A  B"
  shows " P. P' P ReflectL A B"
  using l10_2_existence_spec l10_4_spec by blast

lemma l10_6_existence:
  " P. P' P Reflect A B"
  using l10_2_existence l10_4 by blast

lemma l10_7:
  assumes "P' P Reflect A B" and
    "Q' Q Reflect A B" and
    "P' = Q'"
  shows "P = Q"
  using assms(1) assms(2) assms(3) l10_6_uniqueness by blast

lemma l10_8:
  assumes "P P Reflect A B"
  shows "Col P A B"
  by (metis Col_perm assms col_trivial_2 ex_sym1 l10_6_uniqueness l7_3)

lemma col__refl:
  assumes "Col P A B"
  shows "P P ReflectL A B"
  using ReflectL_def assms col_permutation_1 l7_3_2 by blast

lemma is_image_col_cong:
  assumes "A  B" and
    "P P' Reflect A B" and
    "Col A B X"
  shows "Cong P X P' X"
proof -
  have P1: "P P' ReflectL A B"
    using assms(1) assms(2) is_image_is_image_spec by blast
  obtain M0 where P2: "M0 Midpoint P' P  Col A B M0"
    using P1 ReflectL_def by blast
  have "A B Perp P' P  P' = P"
    using P1 ReflectL_def by auto
  moreover
  {
    assume S1: "A B Perp P' P"
    then have "A  B  P'  P"
      using perp_distinct by blast
    have S2: "M0 = X  Cong P X P' X"
      using P2 cong_4312 midpoint_cong by blast
    {
      assume "M0  X"
      then have "M0 X Perp P' P"
        using P2 S1 assms(3) perp_col2 by blast
      then have "¬ Col A B P  Per P M0 X"
        by (metis Col_perm P2 S1 colx l8_2 midpoint_col midpoint_distinct_1 per_col perp_col1 perp_not_col2 perp_per_1)
      then have "Cong P X P' X"
        using P2 cong_commutativity l7_2 l8_2 per_double_cong by blast
    }
    then have "Cong P X P' X"
      using S2 by blast
  }
  then have "A B Perp P' P  Cong P X P' X" by blast
  moreover
  {
    assume "P = P'"
    then have "Cong P X P' X"
      by (simp add: cong_reflexivity)
  }
  ultimately show ?thesis by blast
qed

lemma is_image_spec_col_cong:
  assumes "P P' ReflectL A B" and
    "Col A B X"
  shows "Cong P X P' X"
  by (metis Col_def Reflect_def assms(1) assms(2) between_trivial col__refl cong_reflexivity is_image_col_cong l10_6_uniqueness_spec)

lemma image_id:
  assumes "A  B" and
    "Col A B T" and
    "T T' Reflect A B"
  shows "T = T'"
  using assms(1) assms(2) assms(3) cong_diff_4 is_image_col_cong by blast

lemma osym_not_col:
  assumes "P P' Reflect A B" and
    "¬ Col A B P"
  shows "¬ Col A B P'"
  using assms(1) assms(2) l10_4 local.image_id not_col_distincts by blast

lemma midpoint_preserves_image:
  assumes "A  B" and
    "Col A B M" and
    "P P' Reflect A B" and
    "M Midpoint P Q" and
    "M Midpoint P' Q'"
  shows "Q Q' Reflect A B"
proof -
  obtain X where P1: "X Midpoint P' P  Col A B X"
    using ReflectL_def assms(1) assms(3) is_image_is_image_spec by blast
  {
    assume S1: "A B Perp P' P"
    obtain Y where S2: "M Midpoint X Y"
      using symmetric_point_construction by blast
    have S3: "Y Midpoint Q Q'"
    proof -
      have R4: "X Midpoint P P'"
        by (simp add: P1 l7_2)
      thus ?thesis
        using assms(4) assms(5) S2 symmetry_preserves_midpoint by blast
    qed
    have S4: "P  P'"
      using S1 perp_not_eq_2 by blast
    then have S5: "Q  Q'"
      using Tarski_neutral_dimensionless.l7_9 Tarski_neutral_dimensionless_axioms assms(4) assms(5) by fastforce
    have S6: "Y Midpoint Q' Q  Col A B Y"
      by (metis P1 S2 S3 assms(2) colx l7_2 midpoint_col midpoint_distinct_1)
    have S7: "A B Perp Q' Q  Q = Q'"
    proof -
      have R3: "Per M Y Q"
      proof -
        have S1: "Y Midpoint Q Q'"
          using S3 by auto
        have "Cong M Q M Q'"
          using assms(1) assms(2) assms(3) assms(4) assms(5) cong_commutativity is_image_col_cong l7_16 l7_3_2 by blast
        thus ?thesis
          using Per_def S1 by blast
      qed
      {
        have "X = Y  (A B Perp Q' Q  Q = Q')"
          by (metis P1 Perp_cases S1 S2 S6 assms(5) l7_3 l7_9_bis)
        {
          assume "X  Y"
          then have "Y PerpAt M Y Y Q"
            using R3 S2 S3 S5 midpoint_distinct_1 per_perp_in by blast
          then have V1: "Y Y Perp Y Q  M Y Perp Y Q"
            by (simp add: perp_in_perp_bis)
          {
            have "Y Y Perp Y Q  A B Perp Q' Q  Q = Q'"
              using perp_not_eq_1 by blast
            {
              assume T1: "M Y Perp Y Q"
              have T2: "Y Q Perp A B"
              proof cases
                assume "A = M"
                thus ?thesis
                  using Perp_cases S6 T1 assms(1) col_permutation_5 perp_col by blast
              next
                assume "A  M"
                thus ?thesis
                  by (smt S6 T1 assms(1) assms(2) col2__eq col_transitivity_2 perp_col0 perp_not_eq_1)
              qed
              have "A B Perp Q' Q  Q = Q'"
                by (metis S3 T2 midpoint_col not_col_distincts perp_col0)
            }
            then have "M Y Perp Y Q  A B Perp Q' Q  Q = Q'" by blast
          }
          then have "A B Perp Q' Q  Q = Q'"
            using V1 perp_distinct by blast
        }
        then have "X  Y  (A B Perp Q' Q  Q = Q')" by blast
      }
      thus ?thesis
        by (metis P1 Perp_cases S1 S2 S6 assms(5) l7_3 l7_9_bis)
    qed
    then have "Q Q' ReflectL A B"
      using ReflectL_def S6 by blast
  }
  then have "A B Perp P' P  Q Q' ReflectL A B" by blast
  moreover
  {
    assume "P = P'"
    then have "Q Q' ReflectL A B"
      by (metis P1 assms(2) assms(4) assms(5) col__refl col_permutation_2 colx midpoint_col midpoint_distinct_3 symmetric_point_uniqueness)
  }
  ultimately show ?thesis
    using ReflectL_def assms(1) assms(3) is_image_is_image_spec by auto
qed

lemma image_in_is_image_spec:
  assumes "M ReflectLAt P P' A B"
  shows "P P' ReflectL A B"
proof -
  have P1: "M Midpoint P' P"
    using ReflectLAt_def assms by blast
  have P2: "Col A B M"
    using ReflectLAt_def assms by blast
  have "A B Perp P' P  P' = P"
    using ReflectLAt_def assms by blast
  thus ?thesis using P1 P2
    using ReflectL_def by blast
qed

lemma image_in_gen_is_image:
  assumes "M ReflectAt P P' A B"
  shows "P P' Reflect A B"
  using ReflectAt_def Reflect_def assms image_in_is_image_spec by auto

lemma image_image_in:
  assumes "P  P'" and
    "P P' ReflectL A B" and
    "Col A B M" and
    "Col P M P'"
  shows "M ReflectLAt P P' A B"
proof -
  obtain M' where P1: "M' Midpoint P' P  Col A B M'"
    using ReflectL_def assms(2) by blast
  have Q1: "P M' Perp A B"
    by (metis Col_cases P1 Perp_perm ReflectL_def assms(1) assms(2) bet_col cong_diff_3 midpoint_bet midpoint_cong not_cong_4321 perp_col1)
  {
    assume R1: "A B Perp P' P"
    have R3: "P  M'"
      using Q1 perp_not_eq_1 by auto
    have R4: "A B Perp P' P"
      by (simp add: R1)
    have R5: "Col P P' M'"
      using P1 midpoint_col not_col_permutation_3 by blast
    have R6: "M' Midpoint P' P"
      by (simp add: P1)
    have R7: "¬ Col A B P"
      using assms(1) assms(2) col__refl col_permutation_2 l10_2_uniqueness_spec l10_4_spec by blast
    have R8: "P  P'"
      by (simp add: assms(1))
    have R9: "Col A B M'"
      by (simp add: P1)
    have R10: "Col A B M"
      by (simp add: assms(3))
    have R11: "Col P P' M'"
      by (simp add: R5)
    have R12: "Col P P' M"
      using Col_perm assms(4) by blast
    have "M = M'"
    proof cases
      assume S1: "A = M'"
      have "Per P M' A"
        by (simp add: S1 l8_5)
      thus ?thesis using l6_21 R8 R9 R10 R11 R12
        using R7 by blast
    next
      assume "A  M'"
      thus ?thesis
        using R10 R12 R5 R7 R8 R9 l6_21 by blast
    qed
    then have "M Midpoint P' P"
      using R6 by blast
  }
  then have Q2: "A B Perp P' P  M Midpoint P' P" by blast
  have Q3: "P' = P  M Midpoint P' P"
    using assms(1) by auto
  have Q4: "A B Perp P' P  P' = P"
    using ReflectL_def assms(2) by auto
  then have "M Midpoint P' P"
    using Q2 Q3 by blast
  thus ?thesis
    by (simp add: ReflectLAt_def Q4 assms(3))
qed

lemma image_in_col:
  assumes "Y ReflectLAt P P' A B"
  shows "Col P P' Y"
  using Col_perm ReflectLAt_def assms midpoint_col by blast

lemma is_image_spec_rev:
  assumes "P P' ReflectL A B"
  shows "P P' ReflectL B A"
proof -
  obtain M0 where P1: "M0 Midpoint P' P  Col A B M0"
    using ReflectL_def assms by blast
  have P2: "Col B A M0"
    using Col_cases P1 by blast
  have "A B Perp P' P  P' = P"
    using ReflectL_def assms by blast
  thus ?thesis
    using P1 P2 Perp_cases ReflectL_def by auto
qed

lemma is_image_rev:
  assumes "P P' Reflect A B"
  shows "P P' Reflect B A"
  using Reflect_def assms is_image_spec_rev by auto

lemma midpoint_preserves_per:
  assumes "Per A B C" and
    "M Midpoint A A1" and
    "M Midpoint B B1" and
    "M Midpoint C C1"
  shows "Per A1 B1 C1"
proof -
  obtain C' where P1: "B Midpoint C C'  Cong A C A C'"
    using Per_def assms(1) by blast
  obtain C1' where P2: "M Midpoint C' C1'"
    using symmetric_point_construction by blast
  thus ?thesis
    by (meson P1 Per_def assms(2) assms(3) assms(4) l7_16 symmetry_preserves_midpoint)
qed

lemma col__image_spec:
  assumes "Col A B X"
  shows "X X ReflectL A B"
  by (simp add: assms col__refl col_permutation_2)

lemma image_triv:
  "A A Reflect A B"
  by (simp add: Reflect_def col__refl col_trivial_1 l7_3_2)

lemma cong_midpoint__image:
  assumes "Cong A X A Y" and
    "B Midpoint X Y"
  shows "Y X Reflect A B"
proof cases
  assume "A = B"
  thus ?thesis
    by (simp add: Reflect_def assms(2))
next
  assume S0: "A  B"
  {
    assume S1: "X  Y"
    then have "X Y Perp A B"
    proof -
      have T1: "B  X"
        using S1 assms(2) midpoint_distinct_1 by blast
      have T2: "B  Y"
        using S1 assms(2) midpoint_not_midpoint by blast
      have "Per A B X"
        using Per_def assms(1) assms(2) by auto
      thus ?thesis
        using S0 S1 T1 T2 assms(2) col_per_perp midpoint_col by auto
    qed
    then have "A B Perp X Y  X = Y"
      using Perp_perm by blast
    then have "Y X Reflect A B"
      using ReflectL_def S0 assms(2) col_trivial_2 is_image_is_image_spec by blast
  }
  then have "X  Y  Y X Reflect A B" by blast
  thus ?thesis
    using assms(2) image_triv is_image_rev l7_3 by blast
qed


lemma col_image_spec__eq:
  assumes "Col A B P" and
    "P P' ReflectL A B"
  shows "P = P'"
  using assms(1) assms(2) col__image_spec l10_2_uniqueness_spec l10_4_spec by blast

lemma image_spec_triv:
  "A A ReflectL B B"
  using col__image_spec not_col_distincts by blast

lemma image_spec__eq:
  assumes "P P' ReflectL A A"
  shows "P = P'"
  using assms col_image_spec__eq not_col_distincts by blast

lemma image__midpoint:
  assumes "P P' Reflect A A"
  shows "A Midpoint P' P"
  using Reflect_def assms by auto

lemma is_image_spec_dec:
  "A B ReflectL C D  ¬ A B ReflectL C D"
  by simp

lemma l10_14:
  assumes "P  P'" and
    "A  B" and
    "P P' Reflect A B"
  shows "A B TS P P'"
proof -
  have P1: "P P' ReflectL A B"
    using assms(2) assms(3) is_image_is_image_spec by blast
  then obtain M0 where "M0 Midpoint P' P  Col A B M0"
    using ReflectL_def by blast
  then have "A B Perp P' P  A B TS P P'"
    by (meson TS_def assms(1) assms(2) assms(3) between_symmetry col_permutation_2 local.image_id midpoint_bet osym_not_col)
  thus ?thesis
    using assms(1) P1 ReflectL_def by blast
qed

lemma l10_15:
  assumes "Col A B C" and
    "¬ Col A B P"
  shows " Q. A B Perp Q C  A B OS P Q"
proof -
  have P1: "A  B"
    using assms(2) col_trivial_1 by auto
  obtain X where P2: "A B TS P X"
    using assms(2) col_permutation_1 l9_10 by blast
  {
    assume Q1: "A = C"
    obtain Q where Q2: " T. A B Perp Q A  Col A B T  Bet X T Q"
      using P1 l8_21 by blast
    then obtain T where "A B Perp Q A  Col A B T  Bet X T Q" by blast
    then have "A B TS Q X"
      by (meson P2 TS_def between_symmetry col_permutation_2 perp_not_col)
    then have Q5: "A B OS P Q"
      using P2 l9_8_1 by blast
    then have " Q. A B Perp Q C  A B OS P Q"
      using Q1 Q2 by blast
  }
  then have P3: "A = C  ( Q. A B Perp Q C  A B OS P Q)" by blast
  {
    assume Q1: "A  C"
    then obtain Q where Q2: " T. C A Perp Q C  Col C A T  Bet X T Q"
      using l8_21 by presburger
    then obtain T where Q3: "C A Perp Q C  Col C A T  Bet X T Q" by blast
    have Q4: "A B Perp Q C"
      using NCol_perm P1 Q2 assms(1) col_trivial_2 perp_col2 by blast
    have "A B TS Q X"
    proof -
      have R1: "¬ Col Q A B"
        using Col_perm P1 Q2 assms(1) col_trivial_2 colx perp_not_col by blast
      have R2: "¬ Col X A B"
        using P2 TS_def by auto
      have R3: "Col T A B"
        by (metis Q1 Q3 assms(1) col_trivial_2 colx not_col_permutation_1)
      have "Bet Q T X"
        using Bet_cases Q3 by blast
      then have " T. Col T A B  Bet Q T X"
        using R3 by blast
      thus ?thesis using R1 R2
        by (simp add: TS_def)
    qed
    then have "A B OS P Q"
      using P2 l9_8_1 by blast
    then have " Q. A B Perp Q C  A B OS P Q"
      using Q4 by blast
  }
  thus ?thesis using P3 by blast
qed

lemma ex_per_cong:
  assumes "A  B" and
    "X  Y" and
    "Col A B C" and
    "¬ Col A B D"
  shows " P. Per P C A  Cong P C X Y  A B OS P D"
proof -
  obtain Q where P1: "A B Perp Q C  A B OS D Q"
    using assms(3) assms(4) l10_15 by blast
  obtain P where P2: "C Out Q P  Cong C P X Y"
    by (metis P1 assms(2) perp_not_eq_2 segment_construction_3)
  have P3: "Per P C A"
    using P1 P2 assms(3) col_trivial_3 l8_16_1 l8_3 out_col by blast
  have "A B OS P D"
    using P1 P2 assms(3) one_side_symmetry os_out_os by blast
  thus ?thesis
    using P2 P3 cong_left_commutativity by blast
qed

lemma exists_cong_per:
  " C. Per A B C  Cong B C X Y"
proof cases
  assume "A = B"
  thus ?thesis
    by (meson Tarski_neutral_dimensionless.l8_5 Tarski_neutral_dimensionless_axioms l8_2 segment_construction)
next
  assume "A  B"
  thus ?thesis
    by (metis Perp_perm bet_col between_trivial l8_16_1 l8_21 segment_construction)
qed

subsubsection "Upper dim 2"

lemma upper_dim_implies_per2__col:
  assumes "upper_dim_axiom"
  shows " A B C X. (Per A X C  X  C  Per B X C)  Col A B X"
proof -
  {
    fix A B C X
    assume "Per A X C  X  C  Per B X C"
    moreover then obtain C' where "X Midpoint C C'  Cong A C A C'"
      using Per_def by blast
    ultimately have "Col A B X"
      by (smt Col_def assms midpoint_cong midpoint_distinct_2 not_cong_2134 per_double_cong upper_dim_axiom_def)
  }
  then show ?thesis by blast
qed

lemma upper_dim_implies_col_perp2__col:
  assumes "upper_dim_axiom"
  shows " A B X Y P. (Col A B P  A B Perp X P  P A Perp Y P)  Col Y X P"
proof -
  {
    fix A B X Y P
    assume H1: "Col A B P  A B Perp X P  P A Perp Y P"
    then have H2: "P  A"
      using perp_not_eq_1 by blast
    have "Col Y X P"
    proof -
      have T1: "Per Y P A"
        using H1 l8_2 perp_per_1 by blast
      moreover have "Per X P A"
        using H1 col_trivial_3 l8_16_1 by blast
      then show ?thesis using T1 H2
        using assms upper_dim_implies_per2__col by blast
    qed
  }
  then show ?thesis by blast
qed

lemma upper_dim_implies_perp2__col:
  assumes "upper_dim_axiom"
  shows " X Y Z A B. (X Y Perp A B  X Z Perp A B)  Col X Y Z"
proof -
  {
    fix X Y Z A B
    assume H1: "X Y Perp A B  X Z Perp A B"
    then have H1A: "X Y Perp A B" by blast
    have H1B: "X Z Perp A B" using H1 by blast
    obtain C where H2: "C PerpAt X Y A B"
      using H1 Perp_def by blast
    obtain C' where H3: "C' PerpAt X Z A B"
      using H1 Perp_def by blast
    have "Col X Y Z"
    proof cases
      assume H2: "Col A B X"
      {
        assume "X = A"
        then have "Col X Y Z" using upper_dim_implies_col_perp2__col
          by (metis H1 H2 Perp_cases assms col_permutation_1)
      }
      then have P1: "X = A  Col X Y Z" by blast
      {
        assume P2: "X  A"
        then have P3: "A B Perp X Y" using perp_sym
          using H1 Perp_perm by blast
        have "Col A B X"
          by (simp add: H2)
        then have P4: "A X Perp X Y" using perp_col
          using P2 P3 by auto
        have P5: "A X Perp X Z"
          by (metis H1 H2 P2 Perp_perm col_trivial_3 perp_col0)
        have P6: "Col Y Z X"
        proof -
          have Q1: "upper_dim_axiom"
            by (simp add: assms)
          have Q2: "Per Y X A"
            using P4 Perp_cases perp_per_2 by blast
          have "Per Z X A"
            by (meson P5 Tarski_neutral_dimensionless.Perp_cases Tarski_neutral_dimensionless_axioms perp_per_2)
          then show ?thesis using Q1 Q2 P2
            using upper_dim_implies_per2__col by blast
        qed
        then have "Col X Y Z"
          using Col_perm by blast
      }
      then show ?thesis
        using P1 by blast
    next
      assume T1: "¬ Col A B X"
      obtain Y0 where Q3: "Y0 PerpAt X Y A B"
        using H1 Perp_def by blast
      obtain Z0 where Q4: "Z0 PerpAt X Z A B"
        using Perp_def H1 by blast
      have Q5: "X Y0 Perp A B"
      proof -
        have R1: "X  Y0"
          using Q3 T1 perp_in_col by blast
        have R2: "X Y Perp A B"
          by (simp add: H1A)
        then show ?thesis using R1
          using Q3 perp_col perp_in_col by blast
      qed
      have "X Z0 Perp A B"
        by (metis H1B Q4 T1 perp_col perp_in_col)
      then have Q7: "Y0 = Z0"
        by (meson Q3 Q4 Q5 T1 Tarski_neutral_dimensionless.Perp_perm Tarski_neutral_dimensionless_axioms l8_18_uniqueness perp_in_col)
      have "Col X Y Z"
      proof -
        have "X  Y0"
          using Q5 perp_distinct by auto
        moreover have "Col X Y0 Y"
          using Q3 not_col_permutation_5 perp_in_col by blast
        moreover have "Col X Y0 Z"
          using Q4 Q7 col_permutation_5 perp_in_col by blast
        ultimately show ?thesis
          using col_transitivity_1 by blast
      qed
      then show ?thesis using l8_18_uniqueness
        by (smt H1 H2 Perp_cases T1 col_trivial_3 perp_col1 perp_in_col perp_not_col)
    qed
  }
  then show ?thesis by blast
qed

lemma upper_dim_implies_not_two_sides_one_side_aux:
  assumes "upper_dim_axiom"
  shows " A B X Y PX. (A  B  PX  A  A B Perp X PX  Col A B PX  ¬ Col X A B  ¬ Col Y A B  ¬ A B TS X Y)  A B OS X Y"
proof -
  {
    fix A B X Y PX
    assume H1: "A  B  PX  A  A B Perp X PX  Col A B PX  ¬ Col X A B  ¬ Col Y A B  ¬ A B TS X Y"
    have H1A: "A  B" using H1 by simp
    have H1B: "PX  A" using H1 by simp
    have H1C: "A B Perp X PX" using H1 by simp
    have H1D: "Col A B PX" using H1 by simp
    have H1E: "¬ Col X A B" using H1 by simp
    have H1F: "¬ Col Y A B" using H1 by simp
    have H1G: "¬ A B TS X Y" using H1 by simp
    have " P T. PX A Perp P PX  Col PX A T  Bet Y T P"
      using H1B l8_21 by blast
    then obtain P T where T1: "PX A Perp P PX  Col PX A T  Bet Y T P"
      by blast
    have J1: "PX A Perp P PX" using T1 by blast
    have J2: "Col PX A T" using T1 by blast
    have J3: "Bet Y T P" using T1 by blast
    have P9: "Col P X PX" using upper_dim_implies_col_perp2__col
      using H1C H1D J1 assms by blast
    have J4: "¬ Col P A B"
      using H1A H1D T1 col_trivial_2 colx not_col_permutation_3 perp_not_col by blast
    have J5: "PX A TS P Y"
    proof -
      have f1: "Col PX A B"
        using H1D not_col_permutation_1 by blast
      then have f2: "Col B PX A"
        using not_col_permutation_1 by blast
      have f3: "p. (T = A  Col p A PX)  ¬ Col p A T"
        by (metis J2 l6_16_1)
      have f4: "Col T PX A"
        using J2 not_col_permutation_1 by blast
      have f5: "p. Col p PX B  ¬ Col p PX A"
        using f2 by (meson H1B l6_16_1)
      have f6: "p. (B = PX  Col p B A)  ¬ Col p B PX"
        using H1D l6_16_1 by blast
      have f7: "p pa. ((B = PX  Col p PX pa)  ¬ Col p PX B)  ¬ Col pa PX A"
        using f5 by (metis l6_16_1)
      have f8: "p. ((T = A  B = PX)  Col p A B)  ¬ Col p A PX"
        using f2 by (metis H1B l6_16_1 not_col_permutation_1)
      have "Col B T PX"
        using f5 f4 not_col_permutation_1 by blast
      then have f9: "p. (T = PX  Col p T B)  ¬ Col p T PX"
        using l6_16_1 by blast
      have "B = PX  ¬ Col Y PX A  ¬ Col P PX A"
        using f1 by (metis (no_types) H1B H1F J4 l6_16_1 not_col_permutation_1)
      then show ?thesis
        using f9 f8 f7 f6 f5 f4 f3 by (metis (no_types) H1B H1F J3 J4 TS_def l9_2 not_col_permutation_1)
    qed
    have J6: "X  PX"
      using H1 perp_not_eq_2 by blast
    have J7: "P  X"
      using H1A H1D H1G J5 col_preserves_two_sides col_trivial_2 not_col_permutation_1 by blast
    have J8: "Bet X PX P  PX Out X P  ¬ Col X PX P"
      using l6_4_2 by blast
    have J9: "PX A TS P X"
      by (metis H1A H1D H1G J5 J6 J8 Out_cases P9 TS_def bet__ts between_symmetry col_permutation_1 col_preserves_two_sides col_trivial_2 l9_5)
    then have "A B OS X Y"
      by (meson H1A H1D J5 col2_os__os col_trivial_2 l9_2 l9_8_1 not_col_permutation_1)
  }
  then show ?thesis by blast
qed

lemma upper_dim_implies_not_two_sides_one_side:
  assumes "upper_dim_axiom"
  shows " A B X Y. (¬ Col X A B  ¬ Col Y A B  ¬ A B TS X Y)  A B OS X Y"
proof -
  {
    fix A B X Y
    assume H1: "¬ Col X A B  ¬ Col Y A B  ¬ A B TS X Y"
    have H1A: "¬ Col X A B" using H1 by simp
    have H1B: "¬ Col Y A B" using H1 by simp
    have H1C: "¬ A B TS X Y" using H1 by simp
    have P1: "A  B"
      using H1A col_trivial_2 by blast
    obtain PX where P2: "Col A B PX  A B Perp X PX"
      using Col_cases H1 l8_18_existence by blast
    have "A B OS X Y"
    proof cases
      assume H5: "PX = A"
      have "B A OS X Y"
      proof -
        have F1: "B A Perp X A"
          using P2 Perp_perm H5 by blast
        have F2: "Col B A A"
          using not_col_distincts by blast
        have F3: "¬ Col X B A"
          using Col_cases H1A by blast
        have F4: "¬ Col Y B A"
          using Col_cases H1B by blast
        have "¬ B A TS X Y"
          using H1C invert_two_sides by blast
        then show ?thesis
          by (metis F1 F3 F4 assms col_trivial_2 upper_dim_implies_not_two_sides_one_side_aux)
      qed
      then show ?thesis
        by (simp add: invert_one_side)
    next
      assume "PX  A"
      then show ?thesis
        using H1 P1 P2 assms upper_dim_implies_not_two_sides_one_side_aux by blast
    qed
  }
  then show ?thesis by blast
qed

lemma upper_dim_implies_not_one_side_two_sides:
  assumes "upper_dim_axiom"
  shows " A B X Y. (¬ Col X A B  ¬ Col Y A B  ¬ A B OS X Y)  A B TS X Y"
  using assms upper_dim_implies_not_two_sides_one_side by blast

lemma upper_dim_implies_one_or_two_sides:
  assumes "upper_dim_axiom"
  shows " A B X Y. (¬ Col X A B  ¬ Col Y A B)  (A B TS X Y  A B OS X Y)"
  using assms upper_dim_implies_not_two_sides_one_side by blast

lemma upper_dim_implies_all_coplanar:
  assumes "upper_dim_axiom"
  shows "all_coplanar_axiom"
  using all_coplanar_axiom_def assms upper_dim_axiom_def by auto

lemma all_coplanar_implies_upper_dim:
  assumes "all_coplanar_axiom"
  shows "upper_dim_axiom"
  using all_coplanar_axiom_def assms upper_dim_axiom_def by auto

lemma all_coplanar_upper_dim:
  shows "all_coplanar_axiom  upper_dim_axiom"
  using all_coplanar_implies_upper_dim upper_dim_implies_all_coplanar by auto

lemma upper_dim_stab:
  shows "¬ ¬ upper_dim_axiom  upper_dim_axiom" by blast

lemma cop__cong_on_bissect:
  assumes "Coplanar A B X P" and
    "M Midpoint A B" and
    "M PerpAt A B P M" and
    "Cong X A X B"
  shows "Col M P X"
proof -
  have P1: "X = M  ¬ Col A B X  M PerpAt X M A B"
    using assms(2) assms(3) assms(4) cong_commutativity cong_perp_or_mid perp_in_distinct by blast
  {
    assume H1: "¬ Col A B X  M PerpAt X M A B"
    then have Q1: "X M Perp A B"
      using perp_in_perp by blast
    have Q2: "A B Perp P M"
      using assms(3) perp_in_perp by blast
    have P2: "Col M A B"
      by (simp add: assms(2) midpoint_col)
    then have "Col M P X" using cop_perp2__col
      by (meson Perp_perm Q1 Q2 assms(1) coplanar_perm_1)
  }
  then show ?thesis
    using P1 not_col_distincts by blast
qed

lemma cong_cop_mid_perp__col:
  assumes "Coplanar A B X P" and
    "Cong A X B X" and
    "M Midpoint A B" and
    "A B Perp P M"
  shows "Col M P X"
proof -
  have "M PerpAt A B P M"
    using Col_perm assms(3) assms(4) bet_col l8_15_1 midpoint_bet by blast
  then show ?thesis
    using assms(1) assms(2) assms(3) cop__cong_on_bissect not_cong_2143 by blast
qed

lemma cop_image_in2__col:
  assumes "Coplanar A B P Q" and
    "M ReflectLAt P P' A B" and
    "M ReflectLAt Q Q' A B"
  shows "Col M P Q"
proof -
  have P1: "P P' ReflectL A B"
    using assms(2) image_in_is_image_spec by auto
  then have P2: "A B Perp P' P  P' = P"
    using ReflectL_def by auto
  have P3: "Q Q' ReflectL A B"
    using assms(3) image_in_is_image_spec by blast
  then have P4: "A B Perp Q' Q  Q' = Q"
    using ReflectL_def by auto
  {
    assume S1: "A B Perp P' P  A B Perp Q' Q"
    {
      assume T1: "A = M"
      have T2: "Per B A P"
        by (metis P1 Perp_perm S1 T1 assms(2) image_in_col is_image_is_image_spec l10_14 perp_col1 perp_distinct perp_per_1 ts_distincts)
      have T3: "Per B A Q"
        by (metis S1 T1 assms(3) image_in_col l8_5 perp_col1 perp_per_1 perp_right_comm)
      have T4: "Coplanar B P Q A"
        using assms(1) ncoplanar_perm_18 by blast
      have T5: "B  A"
        using S1 perp_distinct by blast
      have T6: "Per P A B"
        by (simp add: T2 l8_2)
      have T7: "Per Q A B"
        using Per_cases T3 by blast
      then have "Col P Q A" using T4 T5 T6
        using cop_per2__col by blast
      then have "Col A P Q"
        using not_col_permutation_1 by blast
      then have "Col M P Q"
        using T1 by blast
    }
    then have S2: "A = M  Col M P Q" by blast
    {
      assume D0: "A  M"
      have D1: "Per A M P"
      proof -
        have E1: "M Midpoint P P'"
          using ReflectLAt_def assms(2) l7_2 by blast
        have "Cong P A P' A"
          using P1 col_trivial_3 is_image_spec_col_cong by blast
        then have "Cong A P A P'"
          using Cong_perm by blast
        then show ?thesis
          using E1 Per_def by blast
      qed
      have D2: "Per A M Q"
      proof -
        have E2: "M Midpoint Q Q'"
          using ReflectLAt_def assms(3) l7_2 by blast
        have "Cong A Q A Q'"
          using P3 col_trivial_3 cong_commutativity is_image_spec_col_cong by blast
        then show ?thesis
          using E2 Per_def by blast
      qed
      have "Col P Q M"
      proof -
        have W1: "Coplanar P Q A B"
          using assms(1) ncoplanar_perm_16 by blast
        have W2: "A  B"
          using S1 perp_distinct by blast
        have "Col A B M"
          using ReflectLAt_def assms(2) by blast
        then have "Coplanar P Q A M"
          using W1 W2 col2_cop__cop col_trivial_3 by blast
        then have V1: "Coplanar A P Q M"
          using ncoplanar_perm_8 by blast
        have V3: "Per P M A"
          by (simp add: D1 l8_2)
        have "Per Q M A"
          using D2 Per_perm by blast
        then show ?thesis
          using V1 D0 V3 cop_per2__col by blast
      qed
      then have "Col M P Q"
        using Col_perm by blast
    }
    then have "A  M  Col M P Q" by blast
    then have "Col M P Q"
      using S2 by blast
  }
  then have P5: "(A B Perp P' P  A B Perp Q' Q)  Col M P Q" by blast
  have P6: "(A B Perp P' P  (Q' = Q))  Col M P Q"
    using ReflectLAt_def assms(3) l7_3 not_col_distincts by blast
  have P7: "(P' = P  A B Perp Q' Q)  Col M P Q"
    using ReflectLAt_def assms(2) l7_3 not_col_distincts by blast
  have "(P' = P  Q' = Q)  Col M P Q"
    using ReflectLAt_def assms(3) col_trivial_3 l7_3 by blast
  then show ?thesis
    using P2 P4 P5 P6 P7 by blast
qed

lemma l10_10_spec:
  assumes "P' P ReflectL A B" and
    "Q' Q ReflectL A B"
  shows "Cong P Q P' Q'"
proof cases
  assume "A = B"
  then show ?thesis
    using assms(1) assms(2) cong_reflexivity image_spec__eq by blast
next
  assume H1: "A  B"
  obtain X where P1: "X Midpoint P P'  Col A B X"
    using ReflectL_def assms(1) by blast
  obtain Y where P2: "Y Midpoint Q Q'  Col A B Y"
    using ReflectL_def assms(2) by blast
  obtain Z where P3: "Z Midpoint X Y"
    using midpoint_existence by blast
  have P4: "Col A B Z"
  proof cases
    assume "X = Y"
    then show ?thesis
      by (metis P2 P3 midpoint_distinct_3)
  next
    assume "X  Y"
    then show ?thesis
      by (metis P1 P2 P3 l6_21 midpoint_col not_col_distincts)
  qed
  obtain R where P5: "Z Midpoint P R"
    using symmetric_point_construction by blast
  obtain R' where P6: "Z Midpoint P' R'"
    using symmetric_point_construction by blast
  have P7: "A B Perp P P'  P = P'"
    using ReflectL_def assms(1) by auto
  have P8: "A B Perp Q Q'  Q = Q'"
    using ReflectL_def assms(2) by blast
  {
    assume Q1: "A B Perp P P'  A B Perp Q Q'"
    have Q2: "R R' ReflectL A B"
    proof -
      have "P P' Reflect A B"
        by (simp add: H1 assms(1) is_image_is_image_spec l10_4_spec)
      then have "R R' Reflect A B"
        using H1 P4 P5 P6 midpoint_preserves_image by blast
      then show ?thesis
        using H1 is_image_is_image_spec by blast
    qed
    have Q3: "R  R'"
      using P5 P6 Q1 l7_9 perp_not_eq_2 by blast
    have Q4: "Y Midpoint R R'"
      using P1 P3 P5 P6 symmetry_preserves_midpoint by blast
    have Q5: "Cong Q' R' Q R"
      using P2 Q4 l7_13 by blast
    have Q6: "Cong P' Z P Z"
      using P4 assms(1) is_image_spec_col_cong by auto
    have Q7: "Cong Q' Z Q Z"
      using P4 assms(2) is_image_spec_col_cong by blast
    then have "Cong P Q P' Q'"
    proof -
      have S1: "Cong R Z R' Z"
        using P5 P6 Q6 cong_symmetry l7_16 l7_3_2 by blast
      have "R  Z"
        using Q3 S1 cong_reverse_identity by blast
      then show ?thesis
        by (meson P5 P6 Q5 Q6 Q7 S1 between_symmetry five_segment midpoint_bet not_cong_2143 not_cong_3412)
    qed
  }
  then have P9: "(A B Perp P P'  A B Perp Q Q')  Cong P Q P' Q'" by blast
  have P10: "(A B Perp P P'  Q = Q')  Cong P Q P' Q'"
    using P2 Tarski_neutral_dimensionless.l7_3 Tarski_neutral_dimensionless_axioms assms(1) cong_symmetry is_image_spec_col_cong by fastforce
  have P11: "(P = P'  A B Perp Q Q')  Cong P Q P' Q'"
    using P1 Tarski_neutral_dimensionless.l7_3 Tarski_neutral_dimensionless.not_cong_4321 Tarski_neutral_dimensionless_axioms assms(2) is_image_spec_col_cong by fastforce
  have "(P = P'  Q = Q')  Cong P Q P' Q'"
    using cong_reflexivity by blast
  then show ?thesis
    using P10 P11 P7 P8 P9 by blast
qed

lemma l10_10:
  assumes "P' P Reflect A B" and
    "Q' Q Reflect A B"
  shows "Cong P Q P' Q'"
  using Reflect_def assms(1) assms(2) cong_4321 l10_10_spec l7_13 by auto

lemma image_preserves_bet:
  assumes "A A' ReflectL X Y" and
    "B B' ReflectL X Y" and
    "C C' ReflectL X Y" and
    "Bet A B C"
  shows "Bet A' B' C'"
proof -
  have P3: "A B C Cong3 A' B' C'"
    using Cong3_def assms(1) assms(2) assms(3) l10_10_spec l10_4_spec by blast
  then show ?thesis
    using assms(4) l4_6 by blast
qed

lemma image_gen_preserves_bet:
  assumes "A A' Reflect X Y" and
    "B B' Reflect X Y" and
    "C C' Reflect X Y" and
    "Bet A B C"
  shows "Bet A' B' C'"
proof cases
  assume "X = Y"
  then show ?thesis
    by (metis (full_types) assms(1) assms(2) assms(3) assms(4) image__midpoint l7_15 l7_2)
next
  assume P1: "X  Y"
  then have P2: "A A' ReflectL X Y"
    using assms(1) is_image_is_image_spec by blast
  have P3: "B B' ReflectL X Y"
    using P1 assms(2) is_image_is_image_spec by auto
  have "C C' ReflectL X Y"
    using P1 assms(3) is_image_is_image_spec by blast
  then show ?thesis using image_preserves_bet
    using assms(4) P2 P3 by blast
qed

lemma image_preserves_col:
  assumes "A A' ReflectL X Y" and
    "B B' ReflectL X Y" and
    "C C' ReflectL X Y" and
    "Col A B C"
  shows "Col A' B' C'" using image_preserves_bet
  using Col_def assms(1) assms(2) assms(3) assms(4) by auto

lemma image_gen_preserves_col:
  assumes "A A' Reflect X Y" and
    "B B' Reflect X Y" and
    "C C' Reflect X Y" and
    "Col A B C"
  shows "Col A' B' C'"
  using Col_def assms(1) assms(2) assms(3) assms(4) image_gen_preserves_bet by auto

lemma image_gen_preserves_ncol:
  assumes "A A' Reflect X Y" and
    "B B' Reflect X Y" and
    "C C' Reflect X Y" and
    "¬ Col A B C"
  shows "¬ Col A' B' C'"
  using assms(1) assms(2) assms(3) assms(4)image_gen_preserves_col l10_4 by blast

lemma image_gen_preserves_inter:
  assumes "A A' Reflect X Y" and
    "B B' Reflect X Y" and
    "C C' Reflect X Y" and
    "D D' Reflect X Y" and
    "¬ Col A B C" and
    "C  D" and
    "Col A B I" and
    "Col C D I" and
    "Col A' B' I'" and
    "Col C' D' I'"
  shows "I I' Reflect X Y"
proof -
  obtain I0 where P1: "I I0 Reflect X Y"
    using l10_6_existence by blast
  then show ?thesis
    by (smt Tarski_neutral_dimensionless.image_gen_preserves_col Tarski_neutral_dimensionless_axioms assms(1) assms(10) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) assms(8) assms(9) l10_4 l10_7 l6_21)
qed

lemma intersection_with_image_gen:
  assumes "A A' Reflect X Y" and
    "B B' Reflect X Y" and
    "¬ Col A B A'" and
    "Col A B C" and
    "Col A' B' C"
  shows "Col C X Y"
  by (smt assms(1) assms(2) assms(3) assms(4) assms(5) image_gen_preserves_inter l10_2_uniqueness l10_4 l10_8 not_col_distincts)

lemma image_preserves_midpoint :
  assumes "A A' ReflectL X Y" and
    "B B' ReflectL X Y" and
    "C C' ReflectL X Y" and
    "A Midpoint B C"
  shows "A' Midpoint B' C'"
proof -
  have P1: "Bet B' A' C'" using image_preserves_bet
    using assms(1) assms(2) assms(3) assms(4) midpoint_bet by auto
  have "Cong B' A' A' C'"
    by (metis Cong_perm Tarski_neutral_dimensionless.l10_10_spec Tarski_neutral_dimensionless.l7_13 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) assms(4) cong_transitivity l7_3_2)
  then show ?thesis
    by (simp add: Midpoint_def P1)
qed

lemma image_spec_preserves_per:
  assumes "A A' ReflectL X Y" and
    "B B' ReflectL X Y" and
    "C C' ReflectL X Y" and
    "Per A B C"
  shows "Per A' B' C'"
proof cases
  assume "X = Y"
  then show ?thesis
    using assms(1) assms(2) assms(3) assms(4) image_spec__eq by blast
next
  assume P1: "X  Y"
  obtain C1 where P2: "B Midpoint C C1"
    using symmetric_point_construction by blast
  obtain C1' where P3: "C1 C1' ReflectL X Y"
    by (meson P1 l10_6_existence_spec)
  then have P4: "B' Midpoint C' C1'"
    using P2 assms(2) assms(3) image_preserves_midpoint by blast
  have "Cong A' C' A' C1'"
  proof -
    have Q1: "Cong A' C' A C"
      using assms(1) assms(3) l10_10_spec by auto
    have "Cong A C A' C1'"
      by (metis P2 P3 Tarski_neutral_dimensionless.l10_10_spec Tarski_neutral_dimensionless_axioms assms(1) assms(4) cong_inner_transitivity cong_symmetry per_double_cong)
    then show ?thesis
      using Q1 cong_transitivity by blast
  qed
  then show ?thesis
    using P4 Per_def by blast
qed

lemma image_preserves_per:
  assumes "A A' Reflect X Y" and
    "B B' Reflect X Y"and
    "C C' Reflect X Y" and
    "Per A B C"
  shows "Per A' B' C'"
proof cases
  assume "X = Y"
  then show ?thesis using midpoint_preserves_per
    using assms(1) assms(2) assms(3) assms(4) image__midpoint l7_2  by blast
next
  assume P1: "X  Y"
  have P2: "X  Y  A A' ReflectL X Y"
    using P1 assms(1) is_image_is_image_spec by blast
  have P3: "X  Y  B B' ReflectL X Y"
    using P1 assms(2) is_image_is_image_spec by blast
  have P4: "X  Y  C C' ReflectL X Y"
    using P1 assms(3) is_image_is_image_spec by blast
  then show ?thesis using image_spec_preserves_per
    using P2 P3 assms(4) by blast
qed

lemma l10_12:
  assumes "Per A B C" and
    "Per A' B' C'" and
    "Cong A B A' B'" and
    "Cong B C B' C'"
  shows "Cong A C A' C'"
proof cases
  assume P1: "B = C"
  then have "B' = C'"
    using assms(4) cong_reverse_identity by blast
  then show ?thesis
    using P1 assms(3) by blast
next
  assume P2: "B  C"
  have "Cong A C A' C'"
  proof cases
    assume "A = B"
    then show ?thesis
      using assms(3) assms(4) cong_diff_3 by force
  next
    assume P3: "A  B"
    obtain X where P4: "X Midpoint B B'"
      using midpoint_existence by blast
    obtain A1 where P5: "X Midpoint A' A1"
      using Mid_perm symmetric_point_construction by blast
    obtain C1 where P6: "X Midpoint C' C1"
      using Mid_perm symmetric_point_construction by blast
    have Q1: "A' B' C' Cong3 A1 B C1"
      using Cong3_def P4 P5 P6 l7_13 l7_2 by blast
    have Q2: "Per A1 B C1"
      using assms(2)Q1 l8_10  by blast
    have Q3: "Cong A B A1 B"
      by (metis Cong3_def Q1 Tarski_neutral_dimensionless.cong_symmetry Tarski_neutral_dimensionless_axioms assms(3) cong_inner_transitivity)
    have Q4: "Cong B C B C1"
      by (metis Cong3_def Q1 Tarski_neutral_dimensionless.cong_symmetry Tarski_neutral_dimensionless_axioms assms(4) cong_inner_transitivity)
    obtain Y where P7: "Y Midpoint C C1"
      using midpoint_existence by auto
    then have R1: "C1 C Reflect B Y" using cong_midpoint__image
      using Q4 by blast
    obtain A2 where R2: "A1 A2 Reflect B Y"
      using l10_6_existence by blast
    have R3: "Cong C A2 C1 A1"
      using R1 R2 l10_10 by blast
    have R5: "B B Reflect B Y"
      using image_triv by blast
    have R6: "Per A2 B C" using image_preserves_per
      using Q2 R1 R2 image_triv by blast
    have R7: "Cong A B A2 B"
      using l10_10 Cong_perm Q3 R2 cong_transitivity image_triv by blast
    obtain Z where R7A: "Z Midpoint A A2"
      using midpoint_existence by blast
    have "Cong B A B A2"
      using Cong_perm R7 by blast
    then have T1: "A2 A Reflect B Z" using  R7A cong_midpoint__image
      by blast
    obtain C0 where T2: "B Midpoint C C0"
      using symmetric_point_construction by blast
    have T3: "Cong A C A C0"
      using T2 assms(1) per_double_cong by blast
    have T4: "Cong A2 C A2 C0"
      using R6 T2 per_double_cong by blast
    have T5: "C0 C Reflect B Z"
    proof -
      have "C0 C Reflect Z B"
      proof cases
        assume "A = A2"
        then show ?thesis
          by (metis R7A T2 T3 cong_midpoint__image midpoint_distinct_3)
      next
        assume "A  A2"
        then show ?thesis using l4_17 cong_midpoint__image
          by (metis R7A T2 T3 T4 midpoint_col not_col_permutation_3)
      qed
      then show ?thesis
        using is_image_rev by blast
    qed
    have T6: "Cong A C A2 C0"
      using T1 T5 l10_10 by auto
    have R4: "Cong A C A2 C"
      by (metis T4 T6 Tarski_neutral_dimensionless.cong_symmetry Tarski_neutral_dimensionless_axioms cong_inner_transitivity)
    then have Q5: "Cong A C A1 C1"
      by (meson R3 cong_inner_transitivity not_cong_3421)
    then show ?thesis
      using Cong3_def Q1 Q5 cong_symmetry cong_transitivity by blast
  qed
  then show ?thesis by blast
qed

lemma l10_16:
  assumes "¬ Col A B C" and
    "¬ Col A' B' P" and
    "Cong A B A' B'"
  shows " C'. A B C Cong3 A' B' C'  A' B' OS P C'"
proof cases
  assume "A = B"
  then show ?thesis
    using assms(1) not_col_distincts by auto
next
  assume P1: "A  B"
  obtain X where P2: "Col A B X  A B Perp C X"
    using assms(1) l8_18_existence by blast
  obtain X' where P3: "A B X Cong3 A' B' X'"
    using P2 assms(3) l4_14 by blast
  obtain Q where P4: "A' B' Perp Q X'  A' B' OS P Q"
    using P2 P3 assms(2) l10_15 l4_13 by blast
  obtain C' where P5: "X' Out C' Q  Cong X' C' X C"
    by (metis P2 P4 l6_11_existence perp_distinct)
  have P6: "Cong A C A' C'"
  proof cases
    assume "A = X"
    then show ?thesis
      by (metis Cong3_def P3 P5 cong_4321 cong_commutativity cong_diff_3)
  next
    assume "A  X"
    have P7: "Per A X C"
      using P2 col_trivial_3 l8_16_1 l8_2 by blast
    have P8: "Per A' X' C'"
    proof -
      have "X' PerpAt A' X' X' C'"
      proof -
        have Z1: "A' X' Perp X' C'"
        proof -
          have W1: "X'  C'"
            using P5 out_distinct by blast
          have W2: "X' Q Perp A' B'"
            using P4 Perp_perm by blast
          then have "X' C' Perp A' B'"
            by (metis P5 Perp_perm W1 col_trivial_3 not_col_permutation_5 out_col perp_col2_bis)
          then show ?thesis
            by (metis Cong3_def P2 P3 Perp_perm A  X col_trivial_3 cong_identity l4_13 perp_col2_bis)
        qed
        have Z2: "Col X' A' X'"
          using not_col_distincts by blast
        have "Col X' X' C'"
          by (simp add: col_trivial_1)
        then show ?thesis
          by (simp add: Z1 Z2 l8_14_2_1b_bis)
      qed
      then show ?thesis
        by (simp add: perp_in_per)
    qed
    have P9: "Cong A X A' X'"
      using Cong3_def P3 by auto
    have "Cong X C X' C'"
      using Cong_perm P5 by blast
    then show ?thesis using l10_12
      using P7 P8 P9 by blast
  qed
  have P10: "Cong B C B' C'"
  proof cases
    assume "B = X"
    then show ?thesis
      by (metis Cong3_def P3 P5 cong_4321 cong_commutativity cong_diff_3)
  next
    assume "B  X"
    have Q1: "Per B X C"
      using P2 col_trivial_2 l8_16_1 l8_2 by blast
    have "X' PerpAt B' X' X' C'"
    proof -
      have Q2: "B' X' Perp X' C'"
      proof -
        have R1: "B'  X'"
          using Cong3_def P3 B  X cong_identity by blast
        have "X' C' Perp B' A'"
        proof -
          have S1: "X'  C'"
            using Out_def P5 by blast
          have S2: "X' Q Perp B' A'"
            using P4 Perp_perm by blast
          have "Col X' Q C'"
            using Col_perm P5 out_col by blast
          then show ?thesis
            using S1 S2 perp_col by blast
        qed
        then have R2: "B' A' Perp X' C'"
          using Perp_perm by blast
        have R3: "Col B' A' X'"
          using Col_perm P2 P3 l4_13 by blast
        then show ?thesis
          using R1 R2 perp_col by blast
      qed
      have Q3: "Col X' B' X'"
        by (simp add: col_trivial_3)
      have "Col X' X' C'"
        by (simp add: col_trivial_1)
      then show ?thesis using l8_14_2_1b_bis
        using Q2 Q3 by blast
    qed
    then have Q2: "Per B' X' C'"
      by (simp add: perp_in_per)
    have Q3: "Cong B X B' X'"
      using Cong3_def P3 by blast
    have Q4: "Cong X C X' C'"
      using P5 not_cong_3412 by blast
    then show ?thesis
      using Q1 Q2 Q3 l10_12 by blast
  qed
  have P12: "A' B' OS C' Q  X' Out C' Q  ¬ Col A' B' C'" using l9_19 l4_13
    by (meson P2 P3 P5 one_side_not_col123 out_one_side_1)
  then have P13: "A' B' OS C' Q" using l4_13
    by (meson P2 P3 P4 P5 l6_6 one_side_not_col124 out_one_side_1)
  then show ?thesis
    using Cong3_def P10 P4 P6 assms(3) one_side_symmetry one_side_transitivity by blast
qed

lemma cong_cop_image__col:
  assumes "P  P'" and
    "P P' Reflect A B" and
    "Cong P X P' X" and
    "Coplanar A B P X"
  shows "Col A B X"
proof -
  have P1: "(A  B  P P' ReflectL A B)  (A = B  A Midpoint P' P)"
    by (metis assms(2) image__midpoint is_image_is_image_spec)
  {
    assume Q1: "A  B  P P' ReflectL A B"
    then obtain M where Q2: "M Midpoint P' P  Col A B M"
      using ReflectL_def by blast
    have "Col A B X"
    proof cases
      assume R1: "A = M"
      have R2: "P A Perp A B"
      proof -
        have S1: "P  A"
          using Q2 R1 assms(1) midpoint_distinct_2 by blast
        have S2: "P P' Perp A B"
          using Perp_perm Q1 ReflectL_def assms(1) by blast
        have "Col P P' A"
          using Q2 R1 midpoint_col not_col_permutation_3 by blast
        then show ?thesis using perp_col
          using S1 S2 by blast
      qed
      have R3: "Per P A B"
        by (simp add: R2 perp_comm perp_per_1)
      then have R3A: "Per B A P" using l8_2
        by blast
      have "A Midpoint P P'  Cong X P X P'"
        using Cong_cases Q2 R1 assms(3) l7_2 by blast
      then have R4: "Per X A P"
        using Per_def by blast
      have R5: "Coplanar P B X A"
        using assms(4) ncoplanar_perm_20 by blast
      have "P  A"
        using R2 perp_not_eq_1 by auto
      then show ?thesis using R4 R5 R3A
        using cop_per2__col not_col_permutation_1 by blast
    next
      assume T1: "A  M"
      have T3: "P  M"
        using Q2 assms(1) l7_3_2 sym_preserve_diff by blast
      have T2: "P M Perp M A"
      proof -
        have T4: "P P' Perp M A"
          using Perp_perm Q1 Q2 ReflectL_def T1 assms(1) col_trivial_3 perp_col0 by blast
        have "Col P P' M"
          by (simp add: Col_perm Q2 midpoint_col)
        then show ?thesis using T3 T4 perp_col by blast
      qed
      then have "M P Perp A M"
        using perp_comm by blast
      then have "M PerpAt M P A M"
        using perp_perp_in by blast
      then have "M PerpAt P M M A"
        by (simp add: perp_in_comm)
      then have U1: "Per P M A"
        by (simp add: perp_in_per)
      have U2: "Per X M P" using l7_2 cong_commutativity
        using Per_def Q2 assms(3) by blast
      have "Col A X M"
      proof -
        have W2: "Coplanar P A X M"
          by (meson Q1 Q2 assms(4) col_cop2__cop coplanar_perm_13 ncop_distincts)
        have "Per A M P"
          by (simp add: U1 l8_2)
        then show ?thesis using cop_per2__col
          using U2 T3 W2 by blast
      qed
      then show ?thesis
        using Q2 T1 col2__eq not_col_permutation_4 by blast
    qed
  }
  then have P2: "(A  B  P P' ReflectL A B)  Col A B X" by blast
  have "(A = B  A Midpoint P' P)  Col A B X"
    using col_trivial_1 by blast
  then show ?thesis using P1 P2 by blast
qed

lemma cong_cop_per2_1:
  assumes "A  B" and
    "Per A B X" and
    "Per A B Y" and
    "Cong B X B Y" and
    "Coplanar A B X Y"
  shows "X = Y  B Midpoint X Y"
  by (meson Per_cases assms(1) assms(2) assms(3) assms(4) assms(5) cop_per2__col coplanar_perm_3 l7_20_bis not_col_permutation_5)

lemma cong_cop_per2:
  assumes "A  B" and
    "Per A B X" and
    "Per A B Y" and
    "Cong B X B Y" and
    "Coplanar A B X Y"
  shows "X = Y  X Y ReflectL A B"
proof -
  have "X = Y  B Midpoint X Y"
    using assms(1) assms(2) assms(3) assms(4) assms(5) cong_cop_per2_1 by blast
  then show ?thesis
    by (metis Mid_perm Per_def Reflect_def assms(1) assms(3) cong_midpoint__image symmetric_point_uniqueness)
qed

lemma cong_cop_per2_gen:
  assumes "A  B" and
    "Per A B X" and
    "Per A B Y" and
    "Cong B X B Y" and
    "Coplanar A B X Y"
  shows "X = Y  X Y Reflect A B"
proof -
  have "X = Y  B Midpoint X Y"
    using assms(1) assms(2) assms(3) assms(4) assms(5) cong_cop_per2_1 by blast
  then show ?thesis
    using assms(2) cong_midpoint__image l10_4 per_double_cong by blast
qed

lemma ex_perp_cop:
  assumes "A  B"
  shows " Q. A B Perp Q C  Coplanar A B P Q"
proof -
  {
    assume "Col A B C  Col A B P"
    then have " Q. A B Perp Q C  Coplanar A B P Q"
      using assms ex_ncol_cop l10_15 ncop__ncols by blast
  }
  then have T1: "(Col A B C  Col A B P) 
    ( Q. A B Perp Q C  Coplanar A B P Q)" by blast
  {
    assume "¬Col A B C  Col A B P"
    then have " Q. A B Perp Q C  Coplanar A B P Q"
      by (metis Perp_cases ncop__ncols not_col_distincts perp_exists)
  }
  then have T2: "(¬Col A B C  Col A B P) 
    ( Q. A B Perp Q C  Coplanar A B P Q)" by blast

  {
    assume "Col A B C  ¬Col A B P"
    then have " Q. A B Perp Q C  Coplanar A B P Q"
      using l10_15 os__coplanar by blast
  }
  then have T3: "(Col A B C  ¬Col A B P) 
    ( Q. A B Perp Q C  Coplanar A B P Q)" by blast
  {
    assume "¬Col A B C  ¬Col A B P"
    then have " Q. A B Perp Q C  Coplanar A B P Q"
      using l8_18_existence ncop__ncols perp_right_comm by blast
  }
  then have "(¬Col A B C  ¬Col A B P) 
    ( Q. A B Perp Q C  Coplanar A B P Q)" by blast
  then show ?thesis using T1 T2 T3 by blast
qed

lemma hilbert_s_version_of_pasch_aux:
  assumes "Coplanar A B C P" and
    "¬ Col A I P" and
    "¬ Col B C P" and
    "Bet B I C" and
    "B  I" and
    "I  C" and
    "B  C"
  shows " X. Col I P X  ((Bet A X B  A  X  X  B  A  B)  (Bet A X C  A  X  X  C  A  C))"
proof -
  have T1: "I P TS B C"
    using Col_perm assms(3) assms(4) assms(5) assms(6) bet__ts bet_col col_transitivity_1 by blast
  have T2: "Coplanar A P B I"
    using assms(1) assms(4) bet_cop__cop coplanar_perm_6 ncoplanar_perm_9 by blast
  have T3: "I P TS A B  I P TS A C"
    by (meson T1 T2 TS_def assms(2) cop_nos__ts coplanar_perm_21 l9_2 l9_8_2)
  have T4: "I P TS A B 
( X. Col I P X 
            ((Bet A X B  A  X  X  B  A  B) 
             (Bet A X C  A  X  X  C  A  C)))"
    by (metis TS_def not_col_permutation_2 ts_distincts)
  have "I P TS A C 
( X. Col I P X 
            ((Bet A X B  A  X  X  B  A  B) 
             (Bet A X C  A  X  X  C  A  C)))"
    by (metis TS_def not_col_permutation_2 ts_distincts)

  then show ?thesis using T3 T4 by blast
qed

lemma hilbert_s_version_of_pasch:
  assumes "Coplanar A B C P" and
    "¬ Col C Q P" and
    "¬ Col A B P" and
    "BetS A Q B"
  shows " X. Col P Q X  (BetS A X C  BetS B X C)"
proof -
  obtain X where "Col Q P X 
(Bet C X A  C  X  X  A  C  A 
        Bet C X B  C  X  X  B  C  B)"
    using BetSEq assms(1) assms(2) assms(3) assms(4) coplanar_perm_12 hilbert_s_version_of_pasch_aux by fastforce
  then show ?thesis
    by (metis BetS_def Bet_cases Col_perm)
qed

lemma two_sides_cases:
  assumes "¬ Col PO A B" and
    "PO P OS A B"
  shows  "PO A TS P B  PO B TS P A"
  by (meson assms(1) assms(2) cop_nts__os l9_31 ncoplanar_perm_3 not_col_permutation_4 one_side_not_col124 one_side_symmetry os__coplanar)

lemma not_par_two_sides:
  assumes "C  D" and
    "Col A B I" and
    "Col C D I" and
    "¬ Col A B C"
  shows " X Y. Col C D X  Col C D Y  A B TS X Y"
proof -
  obtain pp :: "'p  'p  'p" where
    f1: "p pa. Bet p pa (pp p pa)  pa  (pp p pa)"
    by (meson point_construction_different)
  then have f2: "p pa pb pc. (Col pc pb p  ¬ Col pc pb (pp p pa))  ¬ Col pc pb pa"
    by (meson Col_def colx)
  have f3: "p pa. Col pa p pa"
    by (meson Col_def between_trivial)
  have f4: "p pa. Col pa p p"
    by (meson Col_def between_trivial)
  have f5: "Col I D C"
    by (meson Col_perm assms(3))
  have f6: "p pa. Col (pp pa p) p pa"
    using f4 f3 f2 by blast
  then have f7: "p pa. Col pa (pp pa p) p"
    by (meson Col_perm)
  then have f8: "p pa pb pc. (pc pb TS p (pp p pa)  Col pc pb p)  ¬ Col pc pb pa"
    using f2 f1 by (meson l9_18)
  have "I = D  Col D (pp D I) C"
    using f7 f5 f3 colx by blast
  then have "I = D  Col C D (pp D I)"
    using Col_perm by blast
  then show ?thesis
    using f8 f6 f3 by (metis Col_perm assms(2) assms(4))
qed

lemma cop_not_par_other_side:
  assumes "C  D" and
    "Col A B I" and
    "Col C D I" and
    "¬ Col A B C" and
    "¬ Col A B P" and
    "Coplanar A B C P"
  shows " Q. Col C D Q  A B TS P Q"
proof -
  obtain X Y where P1:"Col C D X  Col C D Y  A B TS X Y"
    using assms(1) assms(2) assms(3) assms(4) not_par_two_sides by blast
  then have "Coplanar C A B X"
    using Coplanar_def assms(1) assms(2) assms(3) col_transitivity_1 by blast
  then have "Coplanar A B P X"
    using assms(4) assms(6) col_permutation_3 coplanar_trans_1 ncoplanar_perm_2 ncoplanar_perm_6 by blast
  then show ?thesis
    by (meson P1 l9_8_2 TS_def assms(5) cop_nts__os not_col_permutation_2 one_side_symmetry)
qed


lemma cop_not_par_same_side:
  assumes "C  D" and
    "Col A B I" and
    "Col C D I" and
    "¬ Col A B C" and
    "¬ Col A B P" and
    "Coplanar A B C P"
  shows " Q. Col C D Q  A B OS P Q"
proof -
  obtain X Y where P1: "Col C D X  Col C D Y  A B TS X Y"
    using assms(1) assms(2) assms(3) assms(4) not_par_two_sides by blast
  then have "Coplanar C A B X"
    using Coplanar_def assms(1) assms(2) assms(3) col_transitivity_1 by blast
  then have "Coplanar A B P X"
    using assms(4) assms(6) col_permutation_1 coplanar_perm_2 coplanar_trans_1 ncoplanar_perm_14 by blast
  then show ?thesis
    by (meson P1 TS_def assms(5) cop_nts__os l9_2 l9_8_1 not_col_permutation_2)
qed

end

subsubsection "Line reflexivity: 2D"

context Tarski_2D

begin

lemma all_coplanar:
  "Coplanar A B C D"
proof -
  have " A B C P Q. P  Q  Cong A P A Q  Cong B P B Q Cong C P C Q 
(Bet A B C  Bet B C A  Bet C A B)"
    using upper_dim by blast
  then show ?thesis using upper_dim_implies_all_coplanar
    by (smt Tarski_neutral_dimensionless.not_col_permutation_2 Tarski_neutral_dimensionless_axioms all_coplanar_axiom_def all_coplanar_implies_upper_dim coplanar_perm_9 ncop__ncol os__coplanar ts__coplanar upper_dim_implies_not_one_side_two_sides)
qed

lemma per2__col:
  assumes "Per A X C" and
    "X  C" and
    "Per B X C"
  shows "Col A B X"
  using all_coplanar_axiom_def all_coplanar_upper_dim assms(1) assms(2) assms(3) upper_dim upper_dim_implies_per2__col by blast

lemma perp2__col:
  assumes "X Y Perp A B" and
    "X Z Perp A B"
  shows "Col X Y Z"
  by (meson Tarski_neutral_dimensionless.cop_perp2__col Tarski_neutral_dimensionless_axioms all_coplanar assms(1) assms(2))
end

subsection "Angles"

subsubsection "Some generalites"

context Tarski_neutral_dimensionless

begin

lemma l11_3:
  assumes "A B C CongA D E F"
  shows " A' C' D' F'. B Out A' A  B Out C C'  E Out D' D  E Out F F'  A' B C' Cong3 D' E F'"
proof -
  obtain A' C' D' F' where P1: "Bet B A A'  Cong A A' E D  Bet B C C'  Cong C C' E F  Bet E D D'  Cong D D' B A  Bet E F F'  Cong F F' B C  Cong A' C' D' F'" using CongA_def
    using assms by auto
  then have "A' B C' Cong3 D' E F'"
    by (meson Cong3_def between_symmetry l2_11_b not_cong_1243 not_cong_4312)
  thus ?thesis
    by (metis CongA_def P1 assms bet_out l6_6)
qed

lemma l11_aux:
  assumes "B Out A A'" and
    "E Out D D'" and
    "Cong B A' E D'" and
    "Bet B A A0" and
    "Bet E D D0" and
    "Cong A A0 E D" and
    "Cong D D0 B A"
  shows "Cong B A0 E D0  Cong A' A0 D' D0"
proof -
  have P2: "Cong B A0 E D0"
    by (meson Bet_cases assms(4) assms(5) assms(6) assms(7) l2_11_b not_cong_1243 not_cong_4312)
  have P3: "Bet B A A'  Bet B A' A"
    using Out_def assms(1) by auto
  have P4: "Bet E D D'  Bet E D' D"
    using Out_def assms(2) by auto
  have P5: "Bet B A A'  Cong A' A0 D' D0"
    by (smt P2 assms(1) assms(2) assms(3) assms(4) assms(5) bet_out l6_6 l6_7 out_cong_cong out_diff1)
  have P6: "Bet B A' A  Cong A' A0 D' D0"
  proof -
    have "E Out D D0"
      using assms(2) assms(5) bet_out out_diff1 by blast
    thus ?thesis
      by (meson P2 assms(2) assms(3) assms(4) between_exchange4 cong_preserves_bet l4_3_1 l6_6 l6_7)
  qed
  have P7: "Bet E D D'  Cong A' A0 D' D0"
    using P3 P5 P6 by blast
  have "Bet E D' D  Cong A' A0 D' D0"
    using P3 P5 P6 by blast
  thus ?thesis
    using P2 P3 P4 P5 P6 P7 by blast
qed

lemma l11_3_bis:
  assumes " A' C' D' F'. (B Out A' A  B Out C' C  E Out D' D  E Out F' F  A' B C' Cong3 D' E F')"
  shows "A B C CongA