diff --git a/thys/IsaGeoCoq/Tarski_Neutral.thy b/thys/IsaGeoCoq/Tarski_Neutral.thy --- a/thys/IsaGeoCoq/Tarski_Neutral.thy +++ b/thys/IsaGeoCoq/Tarski_Neutral.thy @@ -1,28006 +1,27998 @@ (* 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 D E F" proof - obtain A' C' D' F' where P1: "B Out A' A \ B Out C' C \ E Out D' D \ E Out F' F \ A' B C' Cong3 D' E F'" using assms by blast obtain A0 where P2: "Bet B A A0 \ Cong A A0 E D" using segment_construction by presburger obtain C0 where P3: "Bet B C C0 \ Cong C C0 E F" using segment_construction by presburger obtain D0 where P4: "Bet E D D0 \ Cong D D0 B A" using segment_construction by presburger obtain F0 where P5: "Bet E F F0 \ Cong F F0 B C" using segment_construction by presburger have P6: "A \ B \ C \ B \ D \ E \ F \ E" using P1 out_diff2 by blast have "Cong A0 C0 D0 F0" proof - have Q1: "Cong B A0 E D0 \ Cong A' A0 D' D0" proof - have R1: "B Out A A'" by (simp add: P1 l6_6) have R2: "E Out D D'" by (simp add: P1 l6_6) have "Cong B A' E D'" using Cong3_def P1 cong_commutativity by blast thus ?thesis using l11_aux using P2 P4 R1 R2 by blast qed have Q2: "Cong B C0 E F0 \ Cong C' C0 F' F0" by (smt Cong3_def Out_cases P1 P3 P5 Tarski_neutral_dimensionless.l11_aux Tarski_neutral_dimensionless_axioms) have Q3: "B A' A0 Cong3 E D' D0" by (meson Cong3_def P1 Q1 cong_3_swap) have Q4: "B C' C0 Cong3 E F' F0" using Cong3_def P1 Q2 by blast have "Cong C0 A' F0 D'" proof - have R1: "B C' C0 A' FSC E F' F0 D'" proof - have S1: "Col B C' C0" by (metis (no_types) Col_perm P1 P3 P6 bet_col col_transitivity_1 out_col) have S3: "Cong B A' E D'" using Cong3_def Q3 by blast have "Cong C' A' F' D'" using Cong3_def P1 cong_commutativity by blast thus ?thesis by (simp add: FSC_def S1 Q4 S3) qed have "B \ C'" using P1 out_distinct by blast thus ?thesis using R1 l4_16 by blast qed then have Q6: "B A' A0 C0 FSC E D' D0 F0" by (meson FSC_def P1 P2 P6 Q2 Q3 bet_out l6_7 not_cong_2143 out_col) have "B \ A'" using Out_def P1 by blast thus ?thesis using Q6 l4_16 by blast qed thus ?thesis using P6 P2 P3 P4 P5 CongA_def by auto qed lemma l11_4_1: assumes "A B C CongA D E F" and (*"A \ B" and "C \ B" and "D \ E" and "F \ E" and*) "B Out A' A" and "B Out C' C" and "E Out D' D" and "E Out F' F" and "Cong B A' E D'" and "Cong B C' E F'" shows "Cong A' C' D' F'" proof - obtain A0 C0 D0 F0 where P1: "B Out A0 A \ B Out C C0 \ E Out D0 D \ E Out F F0 \ A0 B C0 Cong3 D0 E F0" using assms(1) l11_3 by blast have P2: "B Out A' A0" using P1 assms(2) l6_6 l6_7 by blast have P3: "E Out D' D0" by (meson P1 assms(4) l6_6 l6_7) have P4: "Cong A' A0 D' D0" proof - have "Cong B A0 E D0" using Cong3_def P1 cong_3_swap by blast thus ?thesis using P2 P3 using assms(6) out_cong_cong by blast qed have P5: "Cong A' C0 D' F0" proof - have P6: "B A0 A' C0 FSC E D0 D' F0" by (meson Cong3_def Cong_perm FSC_def P1 P2 P4 assms(6) not_col_permutation_5 out_col) thus ?thesis using P2 Tarski_neutral_dimensionless.l4_16 Tarski_neutral_dimensionless_axioms out_diff2 by fastforce qed have P6: "B Out C' C0" using P1 assms(3) l6_7 by blast have "E Out F' F0" using P1 assms(5) l6_7 by blast then have "Cong C' C0 F' F0" using Cong3_def P1 P6 assms(7) out_cong_cong by auto then have P9: "B C0 C' A' FSC E F0 F' D'" by (smt Cong3_def Cong_perm FSC_def P1 P5 P6 assms(6) assms(7) not_col_permutation_5 out_col) then have "Cong C' A' F' D'" using P6 Tarski_neutral_dimensionless.l4_16 Tarski_neutral_dimensionless_axioms out_diff2 by fastforce thus ?thesis using Tarski_neutral_dimensionless.not_cong_2143 Tarski_neutral_dimensionless_axioms by fastforce qed lemma l11_4_2: assumes "A \ B" and "C \ B" and "D \ E" and "F \ E" and "\ A' C' D' F'. (B Out A' A \ B Out C' C \ E Out D' D \ E Out F' F \ Cong B A' E D' \ Cong B C' E F' \ Cong A' C' D' F')" shows "A B C CongA D E F" proof - obtain A' where P1: "Bet B A A' \ Cong A A' E D" using segment_construction by fastforce obtain C' where P2: "Bet B C C' \ Cong C C' E F" using segment_construction by fastforce obtain D' where P3: "Bet E D D' \ Cong D D' B A" using segment_construction by fastforce obtain F' where P4: "Bet E F F' \ Cong F F' B C" using segment_construction by fastforce have P5: "Cong A' B D' E" by (meson Bet_cases P1 P3 l2_11_b not_cong_1243 not_cong_4312) have P6: "Cong B C' E F'" by (meson P2 P4 between_symmetry cong_3421 cong_right_commutativity l2_11_b) have "B Out A' A \ B Out C' C \ E Out D' D \ E Out F' F \ A' B C' Cong3 D' E F'" by (metis (no_types, lifting) Cong3_def P1 P2 P3 P4 P5 P6 Tarski_neutral_dimensionless.Out_def Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) assms(4) assms(5) bet_neq12__neq cong_commutativity) thus ?thesis using l11_3_bis by blast qed lemma conga_refl: assumes "A \ B" and "C \ B" shows "A B C CongA A B C" by (meson CongA_def assms(1) assms(2) cong_reflexivity segment_construction) lemma conga_sym: assumes "A B C CongA A' B' C'" shows "A' B' C' CongA A B C" proof - obtain A0 C0 D0 F0 where P1: "Bet B A A0 \ Cong A A0 B' A' \ Bet B C C0 \ Cong C C0 B' C' \ Bet B' A' D0 \ Cong A' D0 B A \ Bet B' C' F0 \ Cong C' F0 B C \ Cong A0 C0 D0 F0" using CongA_def assms by auto thus ?thesis proof - have "\p pa pb pc. Bet B' A' p \ Cong A' p B A \ Bet B' C' pa \ Cong C' pa B C \Bet B A pb \ Cong A pb B' A' \Bet B C pc \ Cong C pc B' C' \ Cong p pa pb pc" by (metis (no_types) Tarski_neutral_dimensionless.cong_symmetry Tarski_neutral_dimensionless_axioms P1) thus ?thesis using CongA_def assms by auto qed qed lemma l11_10: assumes "A B C CongA D E F" and "B Out A' A" and "B Out C' C" and "E Out D' D" and "E Out F' F" shows "A' B C' CongA D' E F'" proof - have P1: "A' \ B" using assms(2) out_distinct by auto have P2: "C' \ B" using Out_def assms(3) by force have P3: "D' \ E" using Out_def assms(4) by blast have P4: "F' \ E" using assms(5) out_diff1 by auto have P5: "\ A'0 C'0 D'0 F'0. (B Out A'0 A' \ B Out C'0 C' \ E Out D'0 D' \ E Out F'0 F' \ Cong B A'0 E D'0 \ Cong B C'0 E F'0) \ Cong A'0 C'0 D'0 F'0" by (meson assms(1) assms(2) assms(3) assms(4) assms(5) l11_4_1 l6_7) thus ?thesis using P1 P2 P3 P4 P5 l11_4_2 by blast qed lemma out2__conga: assumes "B Out A' A" and "B Out C' C" shows "A B C CongA A' B C'" by (smt assms(1) assms(2) between_trivial2 conga_refl l11_10 out2_bet_out out_distinct) lemma cong3_diff: assumes "A \ B" and "A B C Cong3 A' B' C'" shows "A' \ B'" using Cong3_def assms(1) assms(2) cong_diff by blast lemma cong3_diff2: assumes "B \ C" and "A B C Cong3 A' B' C'" shows "B' \ C'" using Cong3_def assms(1) assms(2) cong_diff by blast lemma cong3_conga: assumes "A \ B" and "C \ B" and "A B C Cong3 A' B' C'" shows "A B C CongA A' B' C'" by (metis assms(1) assms(2) assms(3) cong3_diff cong3_diff2 l11_3_bis out_trivial) lemma cong3_conga2: assumes "A B C Cong3 A' B' C'" and "A B C CongA A'' B'' C''" shows "A' B' C' CongA A'' B'' C''" proof - obtain A0 C0 A2 C2 where P1: "Bet B A A0 \ Cong A A0 B'' A'' \ Bet B C C0 \ Cong C C0 B'' C''\ Bet B'' A'' A2 \ Cong A'' A2 B A \ Bet B'' C'' C2 \ Cong C'' C2 B C \ Cong A0 C0 A2 C2" using CongA_def assms(2) by auto obtain A1 where P5: "Bet B' A' A1 \ Cong A' A1 B'' A''" using segment_construction by blast obtain C1 where P6: "Bet B' C' C1 \ Cong C' C1 B'' C''" using segment_construction by blast have P7: "Cong A A0 A' A1" proof - have "Cong B'' A'' A' A1" using P5 using Cong_perm by blast thus ?thesis using Cong_perm P1 cong_inner_transitivity by blast qed have P8: "Cong B A0 B' A1" using Cong3_def P1 P5 P7 assms(1) cong_commutativity l2_11_b by blast have P9: "Cong C C0 C' C1" using P1 P6 cong_inner_transitivity cong_symmetry by blast have P10: "Cong B C0 B' C1" using Cong3_def P1 P6 P9 assms(1) l2_11_b by blast have "B A A0 C FSC B' A' A1 C'" using FSC_def P1 P5 P7 P8 Tarski_neutral_dimensionless.Cong3_def Tarski_neutral_dimensionless_axioms assms(1) bet_col l4_3 by fastforce then have P12: "Cong A0 C A1 C'" using CongA_def assms(2) l4_16 by auto then have "B C C0 A0 FSC B' C' C1 A1" using Cong3_def FSC_def P1 P10 P8 P9 assms(1) bet_col cong_commutativity by auto then have P13: "Cong C0 A0 C1 A1" using l4_16 CongA_def assms(2) by blast have Q2: "Cong A' A1 B'' A''" using P1 P7 cong_inner_transitivity by blast have Q5: "Bet B'' A'' A2" using P1 by blast have Q6: "Cong A'' A2 B' A'" proof - have "Cong B A B' A'" using P1 P7 P8 P5 l4_3 by blast thus ?thesis using P1 cong_transitivity by blast qed have Q7: "Bet B'' C'' C2" using P1 by blast have Q8: "Cong C'' C2 B' C'" proof - have "Cong B C B' C'" using Cong3_def assms(1) by blast thus ?thesis using P1 cong_transitivity by blast qed have R2: "Cong C0 A0 C2 A2" using Cong_cases P1 by blast have "Cong C1 A1 A0 C0" using Cong_cases P13 by blast then have Q9: "Cong A1 C1 A2 C2" using R2 P13 cong_inner_transitivity not_cong_4321 by blast thus ?thesis using CongA_def P5 Q2 P6 Q5 Q6 Q7 Q8 by (metis assms(1) assms(2) cong3_diff cong3_diff2) qed lemma conga_diff1: assumes "A B C CongA A' B' C'" shows "A \ B" using CongA_def assms by blast lemma conga_diff2: assumes "A B C CongA A' B' C'" shows "C \ B" using CongA_def assms by blast lemma conga_diff45: assumes "A B C CongA A' B' C'" shows "A' \ B'" using CongA_def assms by blast lemma conga_diff56: assumes "A B C CongA A' B' C'" shows "C' \ B'" using CongA_def assms by blast lemma conga_trans: assumes "A B C CongA A' B' C'" and "A' B' C' CongA A'' B'' C''" shows "A B C CongA A'' B'' C''" proof - obtain A0 C0 A1 C1 where P1: "Bet B A A0 \ Cong A A0 B' A' \ Bet B C C0 \ Cong C C0 B' C'\ Bet B' A' A1 \ Cong A' A1 B A \ Bet B' C' C1 \ Cong C' C1 B C \ Cong A0 C0 A1 C1" using CongA_def assms(1) by auto have P2: "A''\ B'' \ C'' \ B''" using CongA_def assms(2) by auto have P3: "A1 B' C1 CongA A'' B'' C''" proof - have L2: "B' Out A1 A'" using P1 by (metis Out_def assms(2) bet_neq12__neq conga_diff1) have L3: "B' Out C1 C'" using P1 by (metis Out_def assms(1) bet_neq12__neq conga_diff56) have L4: "B'' Out A'' A''" using P2 out_trivial by auto have "B'' Out C'' C''" by (simp add: P2 out_trivial) thus ?thesis using assms(2) L2 L3 L4 l11_10 by blast qed have L6: "A0 B C0 CongA A' B' C'" by (smt Out_cases P1 Tarski_neutral_dimensionless.conga_diff1 Tarski_neutral_dimensionless.conga_diff2 Tarski_neutral_dimensionless.conga_diff45 Tarski_neutral_dimensionless_axioms assms(1) bet_out conga_diff56 l11_10 l5_3) have L7: "Cong B A0 B' A1" by (meson P1 between_symmetry cong_3421 l2_11_b not_cong_1243) have L8: "Cong B C0 B' C1" using P1 between_symmetry cong_3421 l2_11_b not_cong_1243 by blast have L10: "A0 B C0 Cong3 A1 B' C1" by (simp add: Cong3_def L7 L8 P1 cong_commutativity) then have L11: "A0 B C0 CongA A'' B'' C''" by (meson Tarski_neutral_dimensionless.cong3_conga2 Tarski_neutral_dimensionless_axioms P3 cong_3_sym) thus ?thesis using l11_10 proof - have D2: "B Out A A0" using P1 using CongA_def assms(1) bet_out by auto have D3: "B Out C C0" using P1 using CongA_def assms(1) bet_out by auto have D4: "B'' Out A'' A''" using P2 out_trivial by blast have "B'' Out C'' C''" using P2 out_trivial by auto thus ?thesis using l11_10 L11 D2 D3 D4 by blast qed qed lemma conga_pseudo_refl: assumes "A \ B" and "C \ B" shows "A B C CongA C B A" by (meson CongA_def assms(1) assms(2) between_trivial cong_pseudo_reflexivity segment_construction) lemma conga_trivial_1: assumes "A \ B" and "C \ D" shows "A B A CongA C D C" by (meson CongA_def assms(1) assms(2) cong_trivial_identity segment_construction) lemma l11_13: assumes "A B C CongA D E F" and "Bet A B A'" and "A'\ B" and "Bet D E D'" and "D'\ E" shows "A' B C CongA 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 assms(1) by auto obtain A0 where P2:"Bet B A' A0 \ Cong A' A0 E D'" using segment_construction by blast obtain D0 where P3: "Bet E D' D0 \ Cong D' D0 B A'" using segment_construction by blast have "Cong A0 C'' D0 F''" proof - have L1: "A'' B A0 C'' OFSC D'' E D0 F''" proof - have L2: "Bet A'' B A0" proof - have M1: "Bet A'' A B" using Bet_perm P1 by blast have M2: "Bet A B A0" using P2 assms(2) assms(3) outer_transitivity_between by blast have "A \ B" using CongA_def assms(1) by blast thus ?thesis using M1 M2 outer_transitivity_between2 by blast qed have L3: "Bet D'' E D0" using Bet_perm P1 P2 outer_transitivity_between CongA_def by (metis P3 assms(1) assms(4) assms(5)) have L4: "Cong A'' B D'' E" by (meson P1 between_symmetry cong_3421 cong_left_commutativity l2_11_b) have L5: "Cong B A0 E D0" by (meson P2 P3 between_symmetry cong_3421 cong_right_commutativity l2_11_b) have "Cong B C'' E F''" by (meson P1 between_symmetry cong_3421 cong_right_commutativity l2_11_b) thus ?thesis using P1 L2 L3 L4 L5 by (simp add: OFSC_def) qed have "A'' \ B" using CongA_def P1 assms(1) bet_neq12__neq by fastforce thus ?thesis using L1 five_segment_with_def by auto qed thus ?thesis using CongA_def P1 P2 P3 assms(1) assms(3) assms(5) by auto qed lemma conga_right_comm: assumes "A B C CongA D E F" shows "A B C CongA F E D" by (metis Tarski_neutral_dimensionless.conga_diff45 Tarski_neutral_dimensionless.conga_sym Tarski_neutral_dimensionless.conga_trans Tarski_neutral_dimensionless_axioms assms conga_diff56 conga_pseudo_refl) lemma conga_left_comm: assumes "A B C CongA D E F" shows "C B A CongA D E F" by (meson assms conga_right_comm conga_sym) lemma conga_comm: assumes "A B C CongA D E F" shows "C B A CongA F E D" by (meson Tarski_neutral_dimensionless.conga_left_comm Tarski_neutral_dimensionless.conga_right_comm Tarski_neutral_dimensionless_axioms assms) lemma conga_line: assumes "A \ B" and "B \ C" and "A' \ B'" and "B' \ C'" and "Bet A B C" and "Bet A' B' C'" shows "A B C CongA A' B' C'" by (metis Bet_cases assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) conga_trivial_1 l11_13) lemma l11_14: assumes "Bet A B A'" and "A \ B" and "A' \ B" and "Bet C B C'" and "B \ C" and "B \ C'" shows "A B C CongA A' B C'" by (metis Bet_perm assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) conga_pseudo_refl conga_right_comm l11_13) lemma l11_16: assumes "Per A B C" and "A \ B" and "C \ B" and "Per A' B' C'" and "A'\ B'" and "C'\ B'" shows "A B C CongA A' B' C'" proof - obtain C0 where P1: "Bet B C C0 \ Cong C C0 B' C'" using segment_construction by blast obtain C1 where P2: "Bet B' C' C1 \ Cong C' C1 B C" using segment_construction by blast obtain A0 where P3: "Bet B A A0 \ Cong A A0 B' A'" using segment_construction by blast obtain A1 where P4: "Bet B' A' A1 \ Cong A' A1 B A" using segment_construction by blast have "Cong A0 C0 A1 C1" proof - have Q1: "Per A0 B C0" by (metis P1 P3 Tarski_neutral_dimensionless.l8_3 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) bet_col per_col) have Q2: "Per A1 B' C1" by (metis P2 P4 Tarski_neutral_dimensionless.l8_3 Tarski_neutral_dimensionless_axioms assms(4) assms(5) assms(6) bet_col per_col) have Q3: "Cong A0 B A1 B'" by (meson P3 P4 between_symmetry cong_3421 cong_left_commutativity l2_11_b) have "Cong B C0 B' C1" using P1 P2 between_symmetry cong_3421 l2_11_b not_cong_1243 by blast thus ?thesis using Q1 Q2 Q3 l10_12 by blast qed thus ?thesis using CongA_def P1 P2 P3 P4 assms(2) assms(3) assms(5) assms(6) by auto qed lemma l11_17: assumes "Per A B C" and "A B C CongA A' B' C'" shows "Per A' B' C'" proof - obtain A0 C0 A1 C1 where P1: "Bet B A A0 \ Cong A A0 B' A' \ Bet B C C0 \ Cong C C0 B' C' \ Bet B' A' A1 \ Cong A' A1 B A \ Bet B' C' C1 \ Cong C' C1 B C \ Cong A0 C0 A1 C1" using CongA_def assms(2) by auto have P2: "Per A0 B C0" proof - have Q1: "B \ C" using assms(2) conga_diff2 by blast have Q2: "Per A0 B C" by (metis P1 Tarski_neutral_dimensionless.l8_2 Tarski_neutral_dimensionless_axioms assms(1) assms(2) bet_col conga_diff1 per_col) have "Col B C C0" using P1 bet_col by blast thus ?thesis using Q1 Q2 per_col by blast qed have P3: "Per A1 B' C1" proof - have "A0 B C0 Cong3 A1 B' C1" by (meson Bet_cases Cong3_def P1 l2_11_b not_cong_2134 not_cong_3421) thus ?thesis using P2 l8_10 by blast qed have P4: "B' \ C1" using P1 assms(2) between_identity conga_diff56 by blast have P5: "Per A' B' C1" proof - have P6: "B' \ A1" using P1 assms(2) between_identity conga_diff45 by blast have P7: "Per C1 B' A1" by (simp add: P3 l8_2) have "Col B' A1 A'" using P1 NCol_cases bet_col by blast thus ?thesis using P3 P6 Tarski_neutral_dimensionless.l8_3 Tarski_neutral_dimensionless_axioms by fastforce qed have "Col B' C1 C'" using P1 bet_col col_permutation_5 by blast thus ?thesis using P4 P5 per_col by blast qed lemma l11_18_1: assumes "Bet C B D" and "B \ C" and "B \ D" and "A \ B" and "Per A B C" shows "A B C CongA A B D" by (smt Tarski_neutral_dimensionless.l8_2 Tarski_neutral_dimensionless.l8_5 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) assms(4) assms(5) bet_col col_per2__per l11_16) lemma l11_18_2: assumes "Bet C B D" and "A B C CongA A B D" shows "Per A B C" proof - obtain A0 C0 A1 D0 where P1: "Bet B A A0 \ Cong A A0 B A \ Bet B C C0 \ Cong C C0 B D \ Bet B A A1 \ Cong A A1 B A \ Bet B D D0 \ Cong D D0 B C \ Cong A0 C0 A1 D0" using CongA_def assms(2) by auto have P2: "A0 = A1" by (metis P1 assms(2) conga_diff45 construction_uniqueness) have P3: "Per A0 B C0" proof - have Q1: "Bet C0 B D0" proof - have R1: "Bet C0 C B" using P1 between_symmetry by blast have R2: "Bet C B D0" proof - have S1: "Bet C B D" by (simp add: assms(1)) have S2: "Bet B D D0" using P1 by blast have "B \ D" using assms(2) conga_diff56 by blast thus ?thesis using S1 S2 outer_transitivity_between by blast qed have "C \ B" using assms(2) conga_diff2 by auto thus ?thesis using R1 R2 outer_transitivity_between2 by blast qed have Q2: "Cong C0 B B D0" by (meson P1 between_symmetry cong_3421 l2_11_b not_cong_1243) have "Cong A0 C0 A0 D0" using P1 P2 by blast thus ?thesis using Per_def Q1 Q2 midpoint_def by blast qed have P4: "B \ C0" using P1 assms(2) bet_neq12__neq conga_diff2 by blast have P5: "Per A B C0" by (metis P1 P3 Tarski_neutral_dimensionless.l8_3 Tarski_neutral_dimensionless_axioms assms(2) bet_col bet_col1 bet_neq21__neq col_transitivity_1 conga_diff45) have "Col B C0 C" using P1 using NCol_cases bet_col by blast thus ?thesis using P4 P5 per_col by blast qed lemma cong3_preserves_out: assumes "A Out B C" and "A B C Cong3 A' B' C'" shows "A' Out B' C'" by (meson assms(1) assms(2) col_permutation_4 cong3_symmetry cong_3_swap l4_13 l4_6 not_bet_and_out or_bet_out out_col) lemma l11_21_a: assumes "B Out A C" and "A B C CongA A' B' C'" shows "B' Out A' C'" proof - obtain A0 C0 A1 C1 where P1: "Bet B A A0 \ Cong A A0 B' A' \ Bet B C C0 \ Cong C C0 B' C' \ Bet B' A' A1 \ Cong A' A1 B A \ Bet B' C' C1 \ Cong C' C1 B C \ Cong A0 C0 A1 C1" using CongA_def assms(2) by auto have P2: "B Out A0 C0" by (metis P1 assms(1) bet_out l6_6 l6_7 out_diff1) have P3: "B' Out A1 C1" proof - have "B A0 C0 Cong3 B' A1 C1" by (meson Cong3_def P1 between_symmetry cong_right_commutativity l2_11_b not_cong_4312) thus ?thesis using P2 cong3_preserves_out by blast qed thus ?thesis by (metis P1 assms(2) bet_out conga_diff45 conga_diff56 l6_6 l6_7) qed lemma l11_21_b: assumes "B Out A C" and "B' Out A' C'" shows "A B C CongA A' B' C'" by (smt assms(1) assms(2) between_trivial2 conga_trivial_1 l11_10 out2_bet_out out_distinct) lemma conga_cop__or_out_ts: assumes "Coplanar A B C C'" and "A B C CongA A B C'" shows "B Out C C' \ A B TS C C'" proof - obtain A0 C0 A1 C1 where P1: "Bet B A A0 \ Cong A A0 B A \Bet B C C0 \ Cong C C0 B C' \Bet B A A1 \ Cong A A1 B A \Bet B C' C1 \ Cong C' C1 B C \ Cong A0 C0 A1 C1" using CongA_def assms(2) by auto have P2: "A0 = A1" using P1 by (metis assms(2) conga_diff1 construction_uniqueness) have "B Out C C' \ A B TS C C'" proof cases assume "C0 = C1" thus ?thesis by (metis P1 assms(2) bet2__out conga_diff2 conga_diff56) next assume R1: "C0 \ C1" obtain M where R2: "M Midpoint C0 C1" using midpoint_existence by blast have R3: "Cong B C0 B C1" by (meson Bet_cases P1 l2_11_b not_cong_2134 not_cong_3421) have R3A: "Cong A0 C0 A0 C1" using P1 P2 by blast then have R4: "Per A0 M C0" using R2 using Per_def by blast have R5: "Per B M C0" using Per_def R2 R3 by auto then have R6: "Per B M C1" using R2 l8_4 by blast have R7: "B \ A0" using P1 assms(2) bet_neq12__neq conga_diff45 by blast then have "Cong A C0 A C1" by (meson Col_perm P1 R3 R3A bet_col l4_17) then have R9: "Per A M C0" using Per_def R2 by blast then have R10: "Per A M C1" by (meson R2 Tarski_neutral_dimensionless.l8_4 Tarski_neutral_dimensionless_axioms) have R11: "Col B A M" proof - have S1: "Coplanar C0 B A M" proof - have "Coplanar B A C0 M" proof - have T1: "Coplanar B A C0 C1" proof - have "Coplanar A C0 B C'" proof - have "Coplanar A C' B C0" proof - have U1: "Coplanar A C' B C" by (simp add: assms(1) coplanar_perm_4) have U2: "B \ C" using assms(2) conga_diff2 by blast have "Col B C C0" by (simp add: P1 bet_col) thus ?thesis by (meson Tarski_neutral_dimensionless.col_cop__cop Tarski_neutral_dimensionless_axioms U1 U2) qed thus ?thesis using ncoplanar_perm_5 by blast qed thus ?thesis by (metis P1 Tarski_neutral_dimensionless.col_cop__cop Tarski_neutral_dimensionless_axioms assms(2) bet_col conga_diff56 coplanar_perm_12) qed have "Col C0 C1 M" using Col_perm R2 midpoint_col by blast thus ?thesis using T1 R1 col_cop__cop by blast qed thus ?thesis using ncoplanar_perm_8 by blast qed have "C0 \ M" using R1 R2 midpoint_distinct_1 by blast thus ?thesis using R5 R9 S1 cop_per2__col by blast qed have "B Out C C' \ A B TS C C'" proof cases assume Q1: "B = M" have Q2: "\ Col A B C" by (metis Col_def P1 Q1 R9 assms(2) conga_diff1 conga_diff2 l6_16_1 l8_9 not_bet_and_out out_trivial) have Q3: "\ Col A B C'" by (metis Col_def P1 Q1 R10 assms(2) conga_diff1 conga_diff56 l6_16_1 l8_9 not_bet_and_out out_trivial) have Q4: "Col B A B" by (simp add: col_trivial_3) have "Bet C B C'" proof - have S1: "Bet C1 C' B" using Bet_cases P1 by blast have "Bet C1 B C" proof - have T1: "Bet C0 C B" using Bet_cases P1 by blast have "Bet C0 B C1" by (simp add: Q1 R2 midpoint_bet) thus ?thesis using T1 between_exchange3 between_symmetry by blast qed thus ?thesis using S1 between_exchange3 between_symmetry by blast qed thus ?thesis by (metis Q2 Q3 Q4 bet__ts col_permutation_4 invert_two_sides) next assume L1: "B \ M" have L2: "B M TS C0 C1" proof - have M1: "\ Col C0 B M" by (metis (no_types) Col_perm L1 R1 R2 R5 is_midpoint_id l8_9) have M2: "\ Col C1 B M" using Col_perm L1 R1 R2 R6 l8_9 midpoint_not_midpoint by blast have M3: "Col M B M" using col_trivial_3 by auto have "Bet C0 M C1" by (simp add: R2 midpoint_bet) thus ?thesis using M1 M2 M3 TS_def by blast qed have "A B TS C C'" proof - have W2: "A B TS C C1" proof - have V1: "A B TS C0 C1" using L2 P1 R11 R7 col_two_sides cong_diff invert_two_sides not_col_permutation_5 by blast have "B Out C0 C" using L2 Out_def P1 TS_def assms(2) col_trivial_1 conga_diff2 by auto thus ?thesis using V1 col_trivial_3 l9_5 by blast qed then have W1: "A B TS C' C" proof - have Z1: "A B TS C1 C" by (simp add: W2 l9_2) have Z2: "Col B A B" using not_col_distincts by blast have "B Out C1 C'" using L2 Out_def P1 TS_def assms(2) col_trivial_1 conga_diff56 by auto thus ?thesis using Z1 Z2 l9_5 by blast qed thus ?thesis by (simp add: l9_2) qed thus ?thesis by blast qed thus ?thesis by blast qed thus ?thesis by blast qed lemma conga_os__out: assumes "A B C CongA A B C'" and "A B OS C C'" shows "B Out C C'" using assms(1) assms(2) conga_cop__or_out_ts l9_9 os__coplanar by blast lemma cong2_conga_cong: assumes "A B C CongA 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) cong_4321 l11_3 l11_4_1 not_cong_3412 out_distinct out_trivial) lemma angle_construction_1: assumes "\ Col A B C" and "\ Col A' B' P" shows "\ C'. (A B C CongA A' B' C' \ A' B' OS C' P)" proof - obtain C0 where P1: "Col B A C0 \ B A Perp C C0" using assms(1) col_permutation_4 l8_18_existence by blast have "\ C'. (A B C CongA A' B' C' \ A' B' OS C' P)" proof cases assume P1A: "B = C0" obtain C' where P2: "Per C' B' A' \ Cong C' B' C B \ A' B' OS C' P" by (metis assms(1) assms(2) col_trivial_1 col_trivial_2 ex_per_cong) have P3: "A B C CongA A' B' C'" by (metis P1 P2 Tarski_neutral_dimensionless.l8_2 Tarski_neutral_dimensionless.os_distincts Tarski_neutral_dimensionless_axioms P1A assms(1) l11_16 not_col_distincts perp_per_1) thus ?thesis using P2 by blast next assume P4: "B \ C0" have "\ C'. (A B C CongA A' B' C' \ A' B' OS C' P)" proof cases assume R1: "B Out A C0" obtain C0' where R2: "B' Out A' C0' \ Cong B' C0' B C0" by (metis P4 assms(2) col_trivial_1 segment_construction_3) have "\ C'. Per C' C0' B' \ Cong C' C0' C C0 \ B' C0' OS C' P" proof - have R4: "B' \ C0'" using Out_def R2 by auto have R5: "C \ C0" using P1 perp_distinct by blast have R6: "Col B' C0' C0'" by (simp add: col_trivial_2) have "\ Col B' C0' P" using NCol_cases R2 R4 assms(2) col_transitivity_1 out_col by blast then have "\ C'. Per C' C0' B' \ Cong C' C0' C C0 \ B' C0' OS C' P" using R4 R5 R6 ex_per_cong by blast thus ?thesis by auto qed then obtain C' where R7: "Per C' C0' B' \ Cong C' C0' C C0 \ B' C0' OS C' P" by auto then have R8: "C0 B C Cong3 C0' B' C'" by (meson Cong3_def P1 R2 col_trivial_2 l10_12 l8_16_1 not_col_permutation_2 not_cong_2143 not_cong_4321) have R9: "A B C CongA A' B' C'" proof - have S1: "C0 B C CongA C0' B' C'" by (metis P4 R8 assms(1) cong3_conga not_col_distincts) have S3: "B Out C C" using assms(1) not_col_distincts out_trivial by force have "B' \ C'" using R8 assms(1) cong3_diff2 not_col_distincts by blast then have "B' Out C' C'" using out_trivial by auto thus ?thesis using S1 R1 S3 R2 l11_10 by blast qed have "B' A' OS C' P" proof - have T1: "Col B' C0' A'" by (meson NCol_cases R2 Tarski_neutral_dimensionless.out_col Tarski_neutral_dimensionless_axioms) have "B' \ A'" using assms(2) col_trivial_1 by auto thus ?thesis using T1 R7 col_one_side by blast qed then have "A' B' OS C' P" by (simp add: invert_one_side) thus ?thesis using R9 by blast next assume U1: "\ B Out A C0" then have U2: "Bet A B C0" using NCol_perm P1 or_bet_out by blast obtain C0' where U3: "Bet A' B' C0' \ Cong B' C0' B C0" using segment_construction by blast have U4: "\ C'. Per C' C0' B' \ Cong C' C0' C C0 \ B' C0' OS C' P" proof - have V2: "C \ C0" using Col_cases P1 assms(1) by blast have "B' \ C0'" using P4 U3 cong_diff_3 by blast then have "\ Col B' C0' P" using Col_def U3 assms(2) col_transitivity_1 by blast thus ?thesis using ex_per_cong using V2 not_col_distincts by blast qed then obtain C' where U5: "Per C' C0' B' \ Cong C' C0' C C0 \ B' C0' OS C' P" by blast have U98: "A B C CongA A' B' C'" proof - have X1: "C0 B C Cong3 C0' B' C'" proof - have X2: "Cong C0 B C0' B'" using Cong_cases U3 by blast have X3: "Cong C0 C C0' C'" using U5 not_cong_4321 by blast have "Cong B C B' C'" proof - have Y1: "Per C C0 B" using P1 col_trivial_3 l8_16_1 by blast have "Cong C C0 C' C0'" using U5 not_cong_3412 by blast thus ?thesis using Cong_perm Y1 U5 X2 l10_12 by blast qed thus ?thesis by (simp add: Cong3_def X2 X3) qed have X22: "Bet C0 B A" using U2 between_symmetry by blast have X24: "Bet C0' B' A'" using Bet_cases U3 by blast have "A' \ B'" using assms(2) not_col_distincts by blast thus ?thesis by (metis P4 X1 X22 X24 assms(1) cong3_conga l11_13 not_col_distincts) qed have "A' B' OS C' P" proof - have "B' A' OS C' P" proof - have W1: "Col B' C0' A'" by (simp add: Col_def U3) have "B' \ A'" using assms(2) not_col_distincts by blast thus ?thesis using W1 U5 col_one_side by blast qed thus ?thesis by (simp add: invert_one_side) qed thus ?thesis using U98 by blast qed thus ?thesis by auto qed thus ?thesis by auto qed lemma angle_construction_2: assumes "A \ B" (*and "A \ C"*) and "B \ C" (*and "A' \ B'"*) and "\ Col A' B' P" shows "\ C'. (A B C CongA A' B' C' \ (A' B' OS C' P \ Col A' B' C'))" by (metis Col_def angle_construction_1 assms(1) assms(2) assms(3) col_trivial_3 conga_line l11_21_b or_bet_out out_trivial point_construction_different) lemma ex_conga_ts: assumes "\ Col A B C" and "\ Col A' B' P" shows "\ C'. A B C CongA A' B' C' \ A' B' TS C' P" proof - obtain P' where P1: "A' Midpoint P P'" using symmetric_point_construction by blast have P2: "\ Col A' B' P'" by (metis P1 assms(2) col_transitivity_1 midpoint_col midpoint_distinct_2 not_col_distincts) obtain C' where P3: "A B C CongA A' B' C' \ A' B' OS C' P'" using P2 angle_construction_1 assms(1) by blast have "A' B' TS P' P" using P1 P2 assms(2) bet__ts l9_2 midpoint_bet not_col_distincts by auto thus ?thesis using P3 l9_8_2 one_side_symmetry by blast qed lemma l11_15: assumes "\ Col A B C" and "\ Col D E P" shows "\ F. (A B C CongA D E F \ E D OS F P) \ (\ F1 F2. ((A B C CongA D E F1 \ E D OS F1 P) \ (A B C CongA D E F2 \ E D OS F2 P)) \ E Out F1 F2)" proof - obtain F where P1: "A B C CongA D E F \ D E OS F P" using angle_construction_1 assms(1) assms(2) by blast then have P2: "A B C CongA D E F \ E D OS F P" using invert_one_side by blast have "(\ F1 F2. ((A B C CongA D E F1 \ E D OS F1 P) \ (A B C CongA D E F2 \ E D OS F2 P)) \ E Out F1 F2)" proof - { fix F1 F2 assume P3: "((A B C CongA D E F1 \ E D OS F1 P) \ (A B C CongA D E F2 \ E D OS F2 P))" then have P4: "A B C CongA D E F1" by simp have P5: "E D OS F1 P" using P3 by simp have P6: "A B C CongA D E F2" using P3 by simp have P7: "E D OS F2 P" using P3 by simp have P8: "D E F1 CongA D E F2" using P4 conga_sym P6 conga_trans by blast have "D E OS F1 F2" using P5 P7 invert_one_side one_side_symmetry one_side_transitivity by blast then have "E Out F1 F2" using P8 conga_os__out by (meson P3 conga_sym conga_trans) } thus ?thesis by auto qed thus ?thesis using P2 by blast qed lemma l11_19: assumes "Per A B P1" and "Per A B P2" and "A B OS P1 P2" shows "B Out P1 P2" proof cases assume "Col A B P1" thus ?thesis using assms(3) col123__nos by blast next assume P1: "\ Col A B P1" have "B Out P1 P2" proof cases assume "Col A B P2" thus ?thesis using assms(3) one_side_not_col124 by blast next assume P2: "\ Col A B P2" obtain x where "A B P1 CongA A B x \ B A OS x P2 \ (\ F1 F2. ((A B P1 CongA A B F1 \ B A OS F1 P2) \ (A B P1 CongA A B F2 \ B A OS F2 P2))\ B Out F1 F2)" using P1 P2 l11_15 by blast thus ?thesis by (metis P1 P2 assms(1) assms(2) assms(3) conga_os__out l11_16 not_col_distincts) qed thus ?thesis by simp qed lemma l11_22_bet: assumes "Bet A B C" and "P' B' TS A' C'" and "A B P CongA A' B' P'" and "P B C CongA P' B' C'" shows "Bet A' B' C'" proof - obtain C'' where P1: "Bet A' B' C'' \ Cong B' C'' B C" using segment_construction by blast have P2: "C B P CongA C'' B' P'" by (metis P1 assms(1) assms(3) assms(4) cong_diff_3 conga_diff2 l11_13) have P3: "C'' B' P' CongA C' B' P'" by (meson P2 Tarski_neutral_dimensionless.conga_sym Tarski_neutral_dimensionless_axioms assms(4) conga_comm conga_trans) have P4: "B' Out C' C'' \ P' B' TS C' C''" proof - have P5: "Coplanar P' B' C' C''" by (meson P1 TS_def assms(2) bet__coplanar coplanar_trans_1 ncoplanar_perm_1 ncoplanar_perm_8 ts__coplanar) have "P' B' C' CongA P' B' C''" using P3 conga_comm conga_sym by blast thus ?thesis by (simp add: P5 conga_cop__or_out_ts) qed have P6: "B' Out C' C'' \ Bet A' B' C'" proof - { assume "B' Out C' C''" then have "Bet A' B' C'" using P1 bet_out_out_bet between_exchange4 between_trivial2 col_trivial_3 l6_6 not_bet_out by blast } thus ?thesis by simp qed have "P' B' TS C' C'' \ Bet A' B' C'" proof - { assume P7: "P' B' TS C' C''" then have "Bet A' B' C'" proof cases assume "Col C' B' P'" thus ?thesis using Col_perm TS_def assms(2) by blast next assume Q1: "\ Col C' B' P'" then have Q2: "B' \ P'" using not_col_distincts by blast have Q3: "B' P' TS A' C''" by (metis Col_perm P1 TS_def P7 assms(2) col_trivial_3) have Q4: "B' P' OS C' C''" using P7 Q3 assms(2) invert_two_sides l9_8_1 l9_9 by blast thus ?thesis using P7 invert_one_side l9_9 by blast qed } thus ?thesis by simp qed thus ?thesis using P6 P4 by blast qed lemma l11_22a: assumes "B P TS A C" and "B' P' TS A' C'" and "A B P CongA A' B' P'" and "P B C CongA P' B' C'" shows "A B C CongA A' B' C'" proof - have P1: "A \ B \ A' \ B' \ P \ B \ P' \ B' \ C \ B \ C' \ B'" using assms(3) assms(4) conga_diff1 conga_diff2 conga_diff45 conga_diff56 by auto have P2: "A \ C \ A' \ C'" using assms(1) assms(2) not_two_sides_id by blast obtain A'' where P3: "B' Out A' A'' \ Cong B' A'' B A" using P1 segment_construction_3 by force have P4: "\ Col A B P" using TS_def assms(1) by blast obtain T where P5: "Col T B P \ Bet A T C" using TS_def assms(1) by blast have "A B C CongA A' B' C'" proof cases assume "B = T" thus ?thesis by (metis P1 P5 assms(2) assms(3) assms(4) conga_line invert_two_sides l11_22_bet) next assume P6: "B \ T" have "A B C CongA A' B' C'" proof cases assume P7A: "Bet P B T" obtain T'' where T1: "Bet P' B' T'' \ Cong B' T'' B T" using segment_construction by blast have "\ T''. Col B' P' T'' \ (B' Out P' T'' \ B Out P T) \ Cong B' T'' B T" proof - have T2: "Col B' P' T''" using T1 by (simp add: bet_col col_permutation_4) have "(B' Out P' T'' \ B Out P T) \ Cong B' T'' B T" using P7A T1 not_bet_and_out by blast thus ?thesis using T2 by blast qed then obtain T'' where T3: "Col B' P' T'' \ (B' Out P' T'' \ B Out P T) \ Cong B' T'' B T" by blast then have T4: "B' \ T''" using P6 cong_diff_3 by blast obtain C'' where T5: "Bet A'' T'' C'' \ Cong T'' C'' T C" using segment_construction by blast have T6: "A B T CongA A' B' T''" by (smt Out_cases P5 P6 T3 T4 P7A assms(3) between_symmetry col_permutation_4 conga_comm l11_13 l6_4_1 or_bet_out) then have T7: "A B T CongA A'' B' T''" by (smt P3 P4 P6 T3 Tarski_neutral_dimensionless.l11_10 Tarski_neutral_dimensionless_axioms bet_out col_trivial_3 cong_diff_3 l5_2 l6_6 not_col_permutation_1 or_bet_out) then have T8: "Cong A T A'' T''" using P3 T3 cong2_conga_cong cong_4321 not_cong_3412 by blast have T9: "Cong A C A'' C''" using P5 T5 T8 cong_symmetry l2_11_b by blast have T10: "Cong C B C'' B'" by (smt P3 P4 P5 T3 T5 T8 cong_commutativity cong_symmetry five_segment) have "A B C Cong3 A'' B' C''" using Cong3_def P3 T10 T9 not_cong_2143 not_cong_4321 by blast then have T11: "A B C CongA A'' B' C''" by (simp add: Tarski_neutral_dimensionless.cong3_conga Tarski_neutral_dimensionless_axioms P1) have "C B T Cong3 C'' B' T''" by (simp add: Cong3_def T10 T3 T5 cong_4321 cong_symmetry) then have T12: "C B T CongA C'' B' T''" using P1 P6 cong3_conga by auto have T13: "P B C CongA P' B' C''" proof - have K4: "Bet T B P" using Bet_perm P7A by blast have "Bet T'' B' P'" using Col_perm P7A T3 l6_6 not_bet_and_out or_bet_out by blast thus ?thesis using K4 P1 T12 conga_comm l11_13 by blast qed have T14: "P' B' C' CongA P' B' C''" proof - have "P' B' C' CongA P B C" by (simp add: assms(4) conga_sym) thus ?thesis using T13 conga_trans by blast qed have T15: "B' Out C' C'' \ P' B' TS C' C''" proof - have K7: "Coplanar P' B' C' C''" proof - have K8: "Coplanar A' P' B' C'" using assms(2) coplanar_perm_14 ts__coplanar by blast have K8A: "Coplanar P' C'' B' A''" proof - have "Col P' B' T'' \ Col C'' A'' T''" using Col_def Col_perm T3 T5 by blast then have "Col P' C'' T'' \ Col B' A'' T'' \ Col P' B' T'' \ Col C'' A'' T'' \ Col P' A'' T'' \ Col C'' B' T''" by simp thus ?thesis using Coplanar_def by auto qed then have "Coplanar A' P' B' C''" proof - have K9: "B' \ A''" using P3 out_distinct by blast have "Col B' A'' A'" using Col_perm P3 out_col by blast thus ?thesis using K8A K9 col_cop__cop coplanar_perm_19 by blast qed thus ?thesis by (meson K8 TS_def assms(2) coplanar_perm_7 coplanar_trans_1 ncoplanar_perm_2) qed thus ?thesis by (simp add: T14 K7 conga_cop__or_out_ts) qed have "A B C CongA A' B' C'" proof cases assume "B' Out C' C''" thus ?thesis using P1 P3 T11 l11_10 out_trivial by blast next assume W1: "\ B' Out C' C''" then have W1A: " P' B' TS C' C''" using T15 by simp have W2: "B' P' TS A'' C'" using P3 assms(2) col_trivial_1 l9_5 by blast then have W3: "B' P' OS A'' C''" using T15 W1 invert_two_sides l9_2 l9_8_1 by blast have W4: "B' P' TS A'' C''" proof - have "\ Col A' B' P'" using TS_def assms(2) by auto thus ?thesis using Col_perm T3 T5 W3 one_side_chara by blast qed thus ?thesis using W1A W2 invert_two_sides l9_8_1 l9_9 by blast qed thus ?thesis by simp next assume R1: "\ Bet P B T" then have R2: "B Out P T" using Col_cases P5 l6_4_2 by blast have R2A: "\ T''. Col B' P' T'' \ (B' Out P' T'' \ B Out P T) \ Cong B' T'' B T" proof - obtain T'' where R3: "B' Out P' T'' \ Cong B' T'' B T" using P1 P6 segment_construction_3 by fastforce thus ?thesis using R2 out_col by blast qed then obtain T'' where R4: "Col B' P' T'' \ (B' Out P' T'' \ B Out P T) \ Cong B' T'' B T" by auto have R5: "B' \ T''" using P6 R4 cong_diff_3 by blast obtain C'' where R6: "Bet A'' T'' C'' \ Cong T'' C'' T C" using segment_construction by blast have R7: "A B T CongA A' B' T''" using P1 R2 R4 assms(3) l11_10 l6_6 out_trivial by auto have R8: "A B T CongA A'' B' T''" using P3 P4 R2 R4 assms(3) l11_10 l6_6 not_col_distincts out_trivial by blast have R9: "Cong A T A'' T''" using Cong_cases P3 R4 R8 cong2_conga_cong by blast have R10: "Cong A C A'' C''" using P5 R6 R9 l2_11_b not_cong_3412 by blast have R11: "Cong C B C'' B'" by (smt P3 P4 P5 R4 R6 R9 cong_commutativity cong_symmetry five_segment) have "A B C Cong3 A'' B' C''" by (simp add: Cong3_def P3 R10 R11 cong_4321 cong_commutativity) then have R12: "A B C CongA A'' B' C''" using P1 by (simp add: cong3_conga) have "C B T Cong3 C'' B' T''" using Cong3_def R11 R4 R6 not_cong_3412 not_cong_4321 by blast then have R13: "C B T CongA C'' B' T''" using P1 P6 Tarski_neutral_dimensionless.cong3_conga Tarski_neutral_dimensionless_axioms by fastforce have R13A: "\ Col A' B' P'" using TS_def assms(2) by blast have R14: "B' Out C' C'' \ P' B' TS C' C''" proof - have S1: "Coplanar P' B' C' C''" proof - have T1: "Coplanar A' P' B' C'" using assms(2) ncoplanar_perm_14 ts__coplanar by blast have "Coplanar A' P' B' C''" proof - have U6: "B' \ A''" using P3 out_diff2 by blast have "Coplanar P' C'' B' A''" proof - have "Col P' B' T'' \ Col C'' A'' T''" using Col_def Col_perm R4 R6 by blast thus ?thesis using Coplanar_def by auto qed thus ?thesis by (meson Col_cases P3 U6 col_cop__cop ncoplanar_perm_21 ncoplanar_perm_6 out_col) qed thus ?thesis using NCol_cases R13A T1 coplanar_trans_1 by blast qed have "P' B' C' CongA P' B' C''" proof - have "C B P CongA C'' B' P'" using P1 R12 R13 R2 R4 conga_diff56 l11_10 out_trivial by presburger then have "C' B' P' CongA C'' B' P'" by (meson Tarski_neutral_dimensionless.conga_trans Tarski_neutral_dimensionless_axioms assms(4) conga_comm conga_sym) thus ?thesis by (simp add: conga_comm) qed thus ?thesis by (simp add: S1 conga_cop__or_out_ts) qed have S1: "B Out A A" using P4 not_col_distincts out_trivial by blast have S2: "B Out C C" using TS_def assms(1) not_col_distincts out_trivial by auto have S3: "B' Out A' A''" using P3 by simp have "A B C CongA A' B' C'" proof cases assume "B' Out C' C''" thus ?thesis using S1 S2 S3 using R12 l11_10 by blast next assume "\ B' Out C' C''" then have Z3: "P' B' TS C' C''" using R14 by simp have Q1: "B' P' TS A'' C'" using S3 assms(2) l9_5 not_col_distincts by blast have Q2: "B' P' OS A'' C''" proof - have "B' P' TS C'' C'" proof - have "B' P' TS C' C''" using Z3 using invert_two_sides by blast thus ?thesis by (simp add: l9_2) qed thus ?thesis using Q1 l9_8_1 by blast qed have "B' P' TS A'' C''" using Col_perm Q2 R4 R6 one_side_chara by blast thus ?thesis using Q2 l9_9 by blast qed thus ?thesis using S1 S2 S3 using R12 l11_10 by blast qed thus ?thesis by simp qed thus ?thesis by simp qed lemma l11_22b: assumes "B P OS A C" and "B' P' OS A' C'" and "A B P CongA A' B' P'" and "P B C CongA P' B' C'" shows "A B C CongA A' B' C'" proof - obtain D where P1: "Bet A B D \ Cong B D A B" using segment_construction by blast obtain D' where P2: "Bet A' B' D' \ Cong B' D' A' B'" using segment_construction by blast have P3: "D B P CongA D' B' P'" proof - have Q3: "D \ B" by (metis P1 assms(1) col_trivial_3 cong_diff_3 one_side_not_col124 one_side_symmetry) have Q5: "D' \ B'" by (metis P2 assms(2) col_trivial_3 cong_diff_3 one_side_not_col124 one_side_symmetry) thus ?thesis using assms(3) P1 Q3 P2 l11_13 by blast qed have P5: "D B C CongA D' B' C'" proof - have Q1: "B P TS D C" by (metis P1 assms(1) bet__ts col_trivial_3 cong_diff_3 l9_2 l9_8_2 one_side_not_col124 one_side_symmetry) have "B' P' TS D' C'" by (metis Cong_perm P2 assms(2) bet__ts between_cong between_trivial2 l9_2 l9_8_2 one_side_not_col123 point_construction_different ts_distincts) thus ?thesis using assms(4) Q1 P3 l11_22a by blast qed have P6: "Bet D B A" using Bet_perm P1 by blast have P7: "A \ B" using assms(3) conga_diff1 by auto have P8: "Bet D' B' A'" using Bet_cases P2 by blast have "A' \ B'" using assms(3) conga_diff45 by blast thus ?thesis using P5 P6 P7 P8 l11_13 by blast qed lemma l11_22: assumes "((B P TS A C \ B' P' TS A' C')\(B P OS A C \ B' P' OS A' C'))" and "A B P CongA A' B' P'" and "P B C CongA P' B' C'" shows "A B C CongA A' B' C'" by (meson assms(1) assms(2) assms(3) l11_22a l11_22b) lemma l11_24: assumes "P InAngle A B C" shows "P InAngle C B A" -proof - - obtain pp :: "'p \ 'p \ 'p \ 'p \ 'p" where - "\x0 x1 x2 x3. (\v4. Bet x2 v4 x0 \ (v4 = x1 \ x1 Out v4 x3)) = (Bet x2 (pp x0 x1 x2 x3) x0 \ ((pp x0 x1 x2 x3) = x1 \ x1 Out (pp x0 x1 x2 x3) x3))" - by moura - then have "A \ B \ C \ B \ P \ B \Bet A (pp C B A P) C \ ((pp C B A P) = B \ B Out (pp C B A P) P)" - using InAngle_def assms by presburger - thus ?thesis - by (metis (no_types) InAngle_def between_symmetry) -qed + using Bet_cases InAngle_def assms by auto lemma col_in_angle: assumes "A \ B" and "C \ B" and "P \ B" and "B Out A P \ B Out C P" shows "P InAngle A B C" by (meson InAngle_def assms(1) assms(2) assms(3) assms(4) between_trivial between_trivial2) lemma out321__inangle: assumes "C \ B" and "B Out A P" shows "P InAngle A B C" using assms(1) assms(2) col_in_angle out_distinct by auto lemma inangle1123: assumes "A \ B" and "C \ B" shows "A InAngle A B C" by (simp add: assms(1) assms(2) out321__inangle out_trivial) lemma out341__inangle: assumes "A \ B" and "B Out C P" shows "P InAngle A B C" using assms(1) assms(2) col_in_angle out_distinct by auto lemma inangle3123: assumes "A \ B" and "C \ B" shows "C InAngle A B C" by (simp add: assms(1) assms(2) inangle1123 l11_24) lemma in_angle_two_sides: assumes "\ Col B A P" and "\ Col B C P" and "P InAngle A B C" shows "P B TS A C" by (metis InAngle_def TS_def assms(1) assms(2) assms(3) not_col_distincts not_col_permutation_1 out_col) lemma in_angle_out: assumes "B Out A C" and "P InAngle A B C" shows "B Out A P" by (metis InAngle_def assms(1) assms(2) not_bet_and_out out2_bet_out) lemma col_in_angle_out: assumes "Col B A P" and "\ Bet A B C" and "P InAngle A B C" shows "B Out A P" proof - obtain X where P1: "Bet A X C \ (X = B \ B Out X P)" using InAngle_def assms(3) by auto have "B Out A P" proof cases assume "X = B" thus ?thesis using P1 assms(2) by blast next assume P2: "X \ B" thus ?thesis proof - have f1: "Bet B A P \ A Out B P" by (meson assms(1) l6_4_2) have f2: "B Out X P" using P1 P2 by blast have f3: "(Bet B P A \Bet B A P) \Bet P B A" using f1 by (meson Bet_perm Out_def) have f4: "Bet B P X \Bet P X B" using f2 by (meson Bet_perm Out_def) then have f5: "((Bet B P X \Bet X B A) \Bet B P A) \Bet B A P" using f3 by (meson between_exchange3) have "\p. Bet p X C \ \Bet A p X" using P1 between_exchange3 by blast then have f6: "(P = B \Bet B A P) \Bet B P A" using f5 f3 by (meson Bet_perm P2 assms(2) outer_transitivity_between2) have f7: "Bet C X A" using Bet_perm P1 by blast have "P \ B" using f2 by (simp add: Out_def) moreover { assume "Bet B B C" then have "A \ B" using assms(2) by blast } ultimately have "A \ B" using f7 f4 f1 by (meson Bet_perm Out_def P2 between_exchange3 outer_transitivity_between2) thus ?thesis using f6 f2 by (simp add: Out_def) qed qed thus ?thesis by blast qed lemma l11_25_aux: assumes "P InAngle A B C" and "\ Bet A B C" and "B Out A' A" shows "P InAngle A' B C" proof - have P1: "Bet B A' A \ Bet B A A'" using Out_def assms(3) by auto { assume P2: "Bet B A' A" obtain X where P3: "Bet A X C \ (X = B \ B Out X P)" using InAngle_def assms(1) by auto obtain T where P4: "Bet A' T C \ Bet X T B" using Bet_perm P2 P3 inner_pasch by blast { assume "X = B" then have "P InAngle A' B C" using P3 assms(2) by blast } { assume "B Out X P" then have "P InAngle A' B C" by (metis InAngle_def P4 assms(1) assms(3) bet_out_1 l6_7 out_diff1) } then have "P InAngle A' B C" using P3 \X = B \ P InAngle A' B C\ by blast } { assume Q0: "Bet B A A'" obtain X where Q1: "Bet A X C \ (X = B \ B Out X P)" using InAngle_def assms(1) by auto { assume "X = B" then have "P InAngle A' B C" using Q1 assms(2) by blast } { assume Q2: "B Out X P" obtain T where Q3: "Bet A' T C \ Bet B X T" using Bet_perm Q1 Q0 outer_pasch by blast then have "P InAngle A' B C" by (metis InAngle_def Q0 Q2 assms(1) bet_out l6_6 l6_7 out_diff1) } then have "P InAngle A' B C" using \X = B \ P InAngle A' B C\ Q1 by blast } thus ?thesis using P1 \Bet B A' A \ P InAngle A' B C\ by blast qed lemma l11_25: assumes "P InAngle A B C" and "B Out A' A" and "B Out C' C" and "B Out P' P" shows "P' InAngle A' B C'" proof cases assume "Bet A B C" thus ?thesis by (metis Bet_perm InAngle_def assms(2) assms(3) assms(4) bet_out__bet l6_6 out_distinct) next assume P1: "\ Bet A B C" have P2: "P InAngle A' B C" using P1 assms(1) assms(2) l11_25_aux by blast have P3: "P InAngle A' B C'" proof - have "P InAngle C' B A'" using l11_25_aux using Bet_perm P1 P2 assms(2) assms(3) bet_out__bet l11_24 by blast thus ?thesis using l11_24 by blast qed obtain X where P4: "Bet A' X C' \ (X = B \ B Out X P)" using InAngle_def P3 by auto { assume "X = B" then have "P' InAngle A' B C'" using InAngle_def P3 P4 assms(4) out_diff1 by auto } { assume "B Out X P" then have "P' InAngle A' B C'" proof - have "\p. B Out p P' \ \ B Out p P" by (meson Out_cases assms(4) l6_7) thus ?thesis by (metis (no_types) InAngle_def P3 assms(4) out_diff1) qed } thus ?thesis using InAngle_def P4 assms(2) assms(3) assms(4) out_distinct by auto qed lemma inangle_distincts: assumes "P InAngle A B C" shows "A \ B \ C \ B \ P \ B" using InAngle_def assms by auto lemma segment_construction_0: shows "\ B'. Cong A' B' A B" using segment_construction by blast lemma angle_construction_3: assumes "A \ B" and "C \ B" and "A' \ B'" shows "\ C'. A B C CongA A' B' C'" by (metis angle_construction_2 assms(1) assms(2) assms(3) not_col_exists) lemma l11_28: assumes "A B C Cong3 A' B' C'" and "Col A C D" shows "\ D'. (Cong A D A' D' \ Cong B D B' D' \ Cong C D C' D')" proof cases assume P1: "A = C" have "\ D'. (Cong A D A' D' \ Cong B D B' D' \ Cong C D C' D')" proof cases assume "A = B" thus ?thesis by (metis P1 assms(1) cong3_diff cong3_symmetry cong_3_swap_2 not_cong_3421 segment_construction_0) next assume "A \ B" have "\ D'. (Cong A D A' D' \ Cong B D B' D' \ Cong C D C' D')" proof cases assume "A = D" thus ?thesis using Cong3_def P1 assms(1) cong_trivial_identity by blast next assume "A \ D" have "\ D'. (Cong A D A' D' \ Cong B D B' D' \ Cong C D C' D')" proof cases assume "B = D" thus ?thesis using Cong3_def assms(1) cong_3_swap_2 cong_trivial_identity by blast next assume Q1: "B \ D" obtain D'' where Q2: "B A D CongA B' A' D''" by (metis \A \ B\ \A \ D\ angle_construction_3 assms(1) cong3_diff) obtain D' where Q3: "A' Out D'' D' \ Cong A' D' A D" by (metis Q2 \A \ D\ conga_diff56 segment_construction_3) have Q5: "Cong A D A' D'" using Q3 not_cong_3412 by blast have "B A D CongA B' A' D'" using Q2 Q3 \A \ B\ \A \ D\ conga_diff45 l11_10 l6_6 out_trivial by auto then have "Cong B D B' D'" using Cong3_def Cong_perm Q5 assms(1) cong2_conga_cong by blast thus ?thesis using Cong3_def P1 Q5 assms(1) cong_reverse_identity by blast qed thus ?thesis by simp qed thus ?thesis by simp qed thus ?thesis by simp next assume Z1: "A \ C" have "\ D'. (Cong A D A' D' \ Cong B D B' D' \ Cong C D C' D')" proof cases assume "A = D" thus ?thesis using Cong3_def Cong_perm assms(1) cong_trivial_identity by blast next assume "A \ D" { assume "Bet A C D" obtain D' where W1: "Bet A' C' D' \ Cong C' D' C D" using segment_construction by blast have W2: "Cong A D A' D'" by (meson Cong3_def W1 \Bet A C D\ assms(1) cong_symmetry l2_11_b) have W3: "Cong B D B' D'" proof - have X1: "Cong C D C' D'" using W1 not_cong_3412 by blast have "Cong C B C' B'" using Cong3_def assms(1) cong_commutativity by presburger then have W4: "A C D B OFSC A' C' D' B'" using Cong3_def OFSC_def W1 X1 \Bet A C D\ assms(1) by blast have "Cong D B D' B'" using W4 \A \ C\ five_segment_with_def by blast thus ?thesis using Z1 not_cong_2143 by blast qed have "Cong C D C' D'" by (simp add: W1 cong_symmetry) then have "\ D'. (Cong A D A' D' \ Cong B D B' D' \ Cong C D C' D')" using W2 W3 by blast } { assume W3B: "Bet C D A" then obtain D' where W4A: "Bet A' D' C' \ A D C Cong3 A' D' C'" using Bet_perm Cong3_def assms(1) l4_5 by blast have W5: "Cong A D A' D'" using Cong3_def W4A by blast have "A D C B IFSC A' D' C' B'" by (meson Bet_perm Cong3_def Cong_perm IFSC_def W4A W3B assms(1)) then have "Cong D B D' B'" using l4_2 by blast then have W6: "Cong B D B' D'" using Cong_perm by blast then have "Cong C D C' D'" using Cong3_def W4A not_cong_2143 by blast then have "\ D'. (Cong A D A' D' \ Cong B D B' D' \ Cong C D C' D')" using W5 W6 by blast } { assume W7: "Bet D A C" obtain D' where W7A: "Bet C' A' D' \ Cong A' D' A D" using segment_construction by blast then have W8: "Cong A D A' D'" using Cong_cases by blast have "C A D B OFSC C' A' D' B'" by (meson Bet_perm Cong3_def Cong_perm OFSC_def W7 W7A assms(1)) then have "Cong D B D' B'" using Z1 five_segment_with_def by auto then have w9: "Cong B D B' D'" using Cong_perm by blast have "Cong C D C' D'" proof - have L1: "Bet C A D" using Bet_perm W7 by blast have L2: "Bet C' A' D'" using Bet_perm W7 using W7A by blast have "Cong C A C' A'" using assms(1) using Cong3_def assms(1) not_cong_2143 by blast thus ?thesis using l2_11 using L1 L2 W8 l2_11 by blast qed then have "\ D'. (Cong A D A' D' \ Cong B D B' D' \ Cong C D C' D')" using W8 w9 by blast } thus ?thesis using Bet_cases \Bet A C D \ \D'. Cong A D A' D' \ Cong B D B' D' \ Cong C D C' D'\ \Bet C D A \ \D'. Cong A D A' D' \ Cong B D B' D' \ Cong C D C' D'\ assms(2) third_point by blast qed thus ?thesis by blast qed lemma bet_conga__bet: assumes "Bet A B C" and "A B C CongA A' B' C'" shows "Bet A' B' C'" proof - obtain A0 C0 A1 C1 where P1:" Bet B A A0 \Cong A A0 B' A' \ Bet B C C0 \Cong C C0 B' C' \ Bet B' A' A1 \Cong A' A1 B A \ Bet B' C' C1 \Cong C' C1 B C \ Cong A0 C0 A1 C1" using CongA_def assms(2) by auto have "Bet C B A0" using P1 outer_transitivity_between by (metis assms(1) assms(2) between_symmetry conga_diff1) then have "Bet A0 B C" using Bet_cases by blast then have P2: "Bet A0 B C0" using P1 assms(2) conga_diff2 outer_transitivity_between by blast have P3: "A0 B C0 Cong3 A1 B' C1" proof - have Q1: "Cong A0 B A1 B'" by (meson Bet_cases P1 l2_11_b not_cong_1243 not_cong_4312) have Q3: "Cong B C0 B' C1" using P1 between_symmetry cong_3421 l2_11_b not_cong_1243 by blast thus ?thesis by (simp add: Cong3_def Q1 P1) qed then have P4: "Bet A1 B' C1" using P2 l4_6 by blast then have "Bet A' B' C1" using P1 Bet_cases between_exchange3 by blast thus ?thesis using between_inner_transitivity P1 by blast qed lemma in_angle_one_side: assumes "\ Col A B C" and "\ Col B A P" and "P InAngle A B C" shows "A B OS P C" proof - obtain X where P1: "Bet A X C \ (X = B \ B Out X P)" using InAngle_def assms(3) by auto { assume "X = B" then have "A B OS P C" using P1 assms(1) bet_col by blast } { assume P2: "B Out X P" obtain C' where P2A: "Bet C A C' \ Cong A C' C A" using segment_construction by blast have "A B TS X C'" proof - have Q1: "\ Col X A B" by (metis Col_def P1 assms(1) assms(2) col_transitivity_2 out_col) have Q2 :"\ Col C' A B" by (metis Col_def Cong_perm P2A assms(1) cong_diff l6_16_1) have "\ T. Col T A B \ Bet X T C'" using Bet_cases P1 P2A between_exchange3 col_trivial_1 by blast thus ?thesis by (simp add: Q1 Q2 TS_def) qed then have P3: "A B TS P C'" using P2 col_trivial_3 l9_5 by blast then have "A B TS C C'" by (smt P1 P2 bet_out bet_ts__os between_trivial col123__nos col_trivial_3 invert_two_sides l6_6 l9_2 l9_5) then have "A B OS P C" using OS_def P3 by blast } thus ?thesis using P1 \X = B \ A B OS P C\ by blast qed lemma inangle_one_side: assumes "\ Col A B C" and "\ Col A B P" and "\ Col A B Q" and "P InAngle A B C" and "Q InAngle A B C" shows "A B OS P Q" by (meson assms(1) assms(2) assms(3) assms(4) assms(5) in_angle_one_side not_col_permutation_4 one_side_symmetry one_side_transitivity) lemma inangle_one_side2: assumes "\ Col A B C" and "\ Col A B P" and "\ Col A B Q" and "\ Col C B P" and "\ Col C B Q" and "P InAngle A B C" and "Q InAngle A B C" shows "A B OS P Q \ C B OS P Q" by (meson assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) inangle_one_side l11_24 not_col_permutation_3) lemma col_conga_col: assumes "Col A B C" and "A B C CongA D E F" shows "Col D E F" proof - { assume "Bet A B C" then have "Col D E F" using Col_def assms(2) bet_conga__bet by blast } { assume "Bet B C A" then have "Col D E F" by (meson Col_perm Tarski_neutral_dimensionless.l11_21_a Tarski_neutral_dimensionless_axioms \Bet A B C \ Col D E F\ assms(1) assms(2) or_bet_out out_col) } { assume "Bet C A B" then have "Col D E F" by (meson Col_perm Tarski_neutral_dimensionless.l11_21_a Tarski_neutral_dimensionless_axioms \Bet A B C \ Col D E F\ assms(1) assms(2) or_bet_out out_col) } thus ?thesis using Col_def \Bet A B C \ Col D E F\ \Bet B C A \ Col D E F\ assms(1) by blast qed lemma ncol_conga_ncol: assumes "\ Col A B C" and "A B C CongA D E F" shows "\ Col D E F" using assms(1) assms(2) col_conga_col conga_sym by blast lemma angle_construction_4: assumes "A \ B" and "C \ B" and "A' \ B'" shows "\C'. (A B C CongA A' B' C' \ Coplanar A' B' C' P)" proof cases assume "Col A' B' P" thus ?thesis using angle_construction_3 assms(1) assms(2) assms(3) ncop__ncols by blast next assume "\ Col A' B' P" { assume "Col A B C" then have "\C'. (A B C CongA A' B' C' \ Coplanar A' B' C' P)" by (meson angle_construction_3 assms(1) assms(2) assms(3) col__coplanar col_conga_col) } { assume "\ Col A B C" then obtain C' where "A B C CongA A' B' C' \ A' B' OS C' P" using \\ Col A' B' P\ angle_construction_1 by blast then have "\C'. (A B C CongA A' B' C' \ Coplanar A' B' C' P)" using os__coplanar by blast } thus ?thesis using \Col A B C \ \C'. A B C CongA A' B' C' \ Coplanar A' B' C' P\ by blast qed lemma lea_distincts: assumes "A B C LeA D E F" shows "A\B \ C\B \ D\E \ F\E" by (metis (no_types) LeA_def Tarski_neutral_dimensionless.conga_diff1 Tarski_neutral_dimensionless.conga_diff2 Tarski_neutral_dimensionless_axioms assms inangle_distincts) lemma l11_29_a: assumes "A B C LeA D E F" shows "\ Q. (C InAngle A B Q \ A B Q CongA D E F)" proof - obtain P where P1: "P InAngle D E F \ A B C CongA D E P" using LeA_def assms by blast then have P2: "E \ D \ B \ A \ E \ F \ E \ P \ B \ C" using conga_diff1 conga_diff2 inangle_distincts by blast then have P3: "A \ B \ C \ B" by blast { assume Q1: "Bet A B C" then have Q2: "Bet D E P" by (meson P1 Tarski_neutral_dimensionless.bet_conga__bet Tarski_neutral_dimensionless_axioms) have Q3: "C InAngle A B C" by (simp add: P3 inangle3123) obtain X where Q4: "Bet D X F \ (X = E \ E Out X P)" using InAngle_def P1 by auto have "A B C CongA D E F" proof - { assume R1: "X = E" have R2: "Bet E F P \ Bet E P F" proof - have R3: "D \ E" using P2 by blast have "Bet D E F" using Col_def Col_perm P1 Q2 col_in_angle_out not_bet_and_out by blast have "Bet D E P" using Q2 by blast thus ?thesis using l5_2 using R3 \Bet D E F\ by blast qed then have "A B C CongA D E F" by (smt P1 P2 bet_out l11_10 l6_6 out_trivial) } { assume S1: "E Out X P" have S2: "E Out P F" proof - { assume "Bet E X P" then have "E Out P F" proof - have "Bet E X F" by (meson Bet_perm Q2 Q4 \Bet E X P\ between_exchange3) thus ?thesis by (metis Bet_perm S1 bet2__out between_equality_2 between_trivial2 out2_bet_out out_diff1) qed } { assume "Bet E P X" then have "E Out P F" by (smt Bet_perm Q2 Q4 S1 bet_out_1 between_exchange3 not_bet_and_out outer_transitivity_between2) } thus ?thesis using Out_def S1 \Bet E X P \ E Out P F\ by blast qed then have "A B C CongA D E F" by (metis Bet_perm P2 Q1 Q2 bet_out__bet conga_line) } thus ?thesis using Q4 \X = E \ A B C CongA D E F\ by blast qed then have "\ Q. (C InAngle A B Q \ A B Q CongA D E F)" using conga_diff1 conga_diff2 inangle3123 by blast } { assume "B Out A C" obtain Q where "D E F CongA A B Q" by (metis P2 angle_construction_3) then have "\ Q. (C InAngle A B Q \ A B Q CongA D E F)" by (metis Tarski_neutral_dimensionless.conga_comm Tarski_neutral_dimensionless_axioms \B Out A C\ conga_diff1 conga_sym out321__inangle) } { assume ZZ: "\ Col A B C" have Z1: "D \ E" using P2 by blast have Z2: "F \ E" using P2 by blast have Z3: "Bet D E F \ E Out D F \ \ Col D E F" using not_bet_out by blast { assume "Bet D E F" obtain Q where Z4: "Bet A B Q \ Cong B Q E F" using segment_construction by blast then have "\ Q. (C InAngle A B Q \ A B Q CongA D E F)" by (metis InAngle_def P3 Z1 Z2 \Bet D E F\ conga_line point_construction_different) } { assume "E Out D F" then have Z5: "E Out D P" using P1 in_angle_out by blast have "D E P CongA A B C" by (simp add: P1 conga_sym) then have Z6: "B Out A C" using l11_21_a Z5 by blast then have "\ Q. (C InAngle A B Q \ A B Q CongA D E F)" using \B Out A C \ \Q. C InAngle A B Q \ A B Q CongA D E F\ by blast } { assume W1: "\ Col D E F" obtain Q where W2: "D E F CongA A B Q \ A B OS Q C" using W1 ZZ angle_construction_1 by blast obtain DD where W3: "E Out D DD \ Cong E DD B A" using P3 Z1 segment_construction_3 by force obtain FF where W4: "E Out F FF \ Cong E FF B Q" by (metis P2 W2 conga_diff56 segment_construction_3) then have W5: "P InAngle DD E FF" by (smt Out_cases P1 P2 W3 l11_25 out_trivial) obtain X where W6: "Bet DD X FF \ (X = E \ E Out X P)" using InAngle_def W5 by presburger { assume W7: "X = E" have W8: "Bet D E F" proof - have W10: "E Out DD D" by (simp add: W3 l6_6) have "E Out FF F" by (simp add: W4 l6_6) thus ?thesis using W6 W7 W10 bet_out_out_bet by blast qed then have "\ Q. (C InAngle A B Q \ A B Q CongA D E F)" using \Bet D E F \ \Q. C InAngle A B Q \ A B Q CongA D E F\ by blast } { assume V1: "E Out X P" have "B \ C \ E \ X" using P3 V1 out_diff1 by blast then obtain CC where V2: "B Out C CC \ Cong B CC E X" using segment_construction_3 by blast then have V3: "A B CC CongA DD E X" by (smt P1 P2 V1 W3 l11_10 l6_6 out_trivial) have V4: "Cong A CC DD X" proof - have "Cong A B DD E" using W3 not_cong_4321 by blast thus ?thesis using V2 V3 cong2_conga_cong by blast qed have V5: "A B Q CongA DD E FF" proof - have U1: "D E F CongA A B Q" by (simp add: W2) then have U1A: "A B Q CongA D E F" by (simp add: conga_sym) have U2: "B Out A A" by (simp add: P3 out_trivial) have U3: "B Out Q Q" using W2 conga_diff56 out_trivial by blast have U4: "E Out DD D" using W3 l6_6 by blast have "E Out FF F" by (simp add: W4 l6_6) thus ?thesis using l11_10 using U1A U2 U3 U4 by blast qed then have V6: "Cong A Q DD FF" using Cong_perm W3 W4 cong2_conga_cong by blast have "CC B Q CongA X E FF" proof - have U1: "B A OS CC Q" by (metis (no_types) V2 W2 col124__nos invert_one_side one_side_symmetry one_side_transitivity out_one_side) have U2: "E DD OS X FF" proof - have "\ Col DD E FF" by (metis Col_perm OS_def TS_def U1 V5 ncol_conga_ncol) then have "\ Col E DD X" by (metis Col_def V2 V4 W6 ZZ cong_identity l6_16_1 os_distincts out_one_side) then have "DD E OS X FF" by (metis Col_perm W6 bet_out not_col_distincts one_side_reflexivity out_out_one_side) thus ?thesis by (simp add: invert_one_side) qed have "CC B A CongA X E DD" by (simp add: V3 conga_comm) thus ?thesis using U1 U2 V5 l11_22b by blast qed then have V8: "Cong CC Q X FF" using V2 W4 cong2_conga_cong cong_commutativity not_cong_3412 by blast have V9: "CC InAngle A B Q" proof - have T2: "Q \ B" using W2 conga_diff56 by blast have T3: "CC \ B" using V2 out_distinct by blast have "Bet A CC Q" proof - have T4: "DD X FF Cong3 A CC Q" using Cong3_def V4 V6 V8 not_cong_3412 by blast thus ?thesis using W6 l4_6 by blast qed then have "\ X0. Bet A X0 Q \ (X0 = B \ B Out X0 CC)" using out_trivial by blast thus ?thesis by (simp add: InAngle_def P3 T2 T3) qed then have "C InAngle A B Q" using V2 inangle_distincts l11_25 out_trivial by blast then have "\ Q. (C InAngle A B Q \ A B Q CongA D E F)" using W2 conga_sym by blast } then have "\ Q. (C InAngle A B Q \ A B Q CongA D E F)" using W6 \X = E \ \Q. C InAngle A B Q \ A B Q CongA D E F\ by blast } then have "\ Q. (C InAngle A B Q \ A B Q CongA D E F)" using Z3 \E Out D F \ \Q. C InAngle A B Q \ A B Q CongA D E F\ \Bet D E F \ \Q. C InAngle A B Q \ A B Q CongA D E F\ by blast } thus ?thesis using \B Out A C \ \Q. C InAngle A B Q \ A B Q CongA D E F\ \Bet A B C \ \Q. C InAngle A B Q \ A B Q CongA D E F\ not_bet_out by blast qed lemma in_angle_line: assumes "P \ B" and "A \ B" and "C \ B" and "Bet A B C" shows "P InAngle A B C" using InAngle_def assms(1) assms(2) assms(3) assms(4) by auto lemma l11_29_b: assumes "\ Q. (C InAngle A B Q \ A B Q CongA D E F)" shows "A B C LeA D E F" proof - obtain Q where P1: "C InAngle A B Q \ A B Q CongA D E F" using assms by blast obtain X where P2: "Bet A X Q \ (X = B \ B Out X C)" using InAngle_def P1 by auto { assume P2A: "X = B" obtain P where P3: "A B C CongA D E P" using angle_construction_3 assms conga_diff45 inangle_distincts by fastforce have "P InAngle D E F" proof - have O1: "Bet D E F" by (metis (no_types) P1 P2 Tarski_neutral_dimensionless.bet_conga__bet Tarski_neutral_dimensionless_axioms P2A) have O2: "P \ E" using P3 conga_diff56 by auto have O3: "D \ E" using P3 conga_diff45 by auto have "F \ E" using P1 conga_diff56 by blast thus ?thesis using in_angle_line by (simp add: O1 O2 O3) qed then have "A B C LeA D E F" using LeA_def P3 by blast } { assume G1: "B Out X C" obtain DD where G2: "E Out D DD \ Cong E DD B A" by (metis assms conga_diff1 conga_diff45 segment_construction_3) have G3: "D \ E \ DD \ E" using G2 out_diff1 out_diff2 by blast obtain FF where G3G: "E Out F FF \ Cong E FF B Q" by (metis P1 conga_diff56 inangle_distincts segment_construction_3) then have G3A: "F \ E" using out_diff1 by blast have G3B: "FF \ E" using G3G out_distinct by blast have G4: "Bet A B C \ B Out A C \ \ Col A B C" using not_bet_out by blast { assume G5: "Bet A B C" have G6: "F InAngle D E F" by (simp add: G3 G3A inangle3123) have "A B C CongA D E F" by (smt Bet_perm G3 G3A G5 Out_def P1 P2 bet_conga__bet between_exchange3 conga_line inangle_distincts outer_transitivity_between2) then have "A B C LeA D E F" using G6 LeA_def by blast } { assume G8: "B Out A C" have G9: "D InAngle D E F" by (simp add: G3 G3A inangle1123) have "A B C CongA D E D" by (simp add: G3 G8 l11_21_b out_trivial) then have "A B C LeA D E F" using G9 LeA_def by blast } { assume R1: "\ Col A B C" have R2: "Bet A B Q \ B Out A Q \ \ Col A B Q" using not_bet_out by blast { assume R3: "Bet A B Q" obtain P where R4: "A B C CongA D E P" by (metis G3 LeA_def \Bet A B C \ A B C LeA D E F\ angle_construction_3 not_bet_distincts) have R5: "P InAngle D E F" proof - have R6: "P \ E" using R4 conga_diff56 by auto have "Bet D E F" by (metis (no_types) P1 R3 Tarski_neutral_dimensionless.bet_conga__bet Tarski_neutral_dimensionless_axioms) thus ?thesis by (simp add: R6 G3 G3A in_angle_line) qed then have "A B C LeA D E F" using R4 R5 LeA_def by blast } { assume S1: "B Out A Q" have S2: "B Out A C" using G1 P2 S1 l6_7 out_bet_out_1 by blast have S3: "Col A B C" by (simp add: Col_perm S2 out_col) then have "A B C LeA D E F" using R1 by blast } { assume S3B: "\ Col A B Q" obtain P where S4: "A B C CongA D E P \ D E OS P F" by (meson P1 R1 Tarski_neutral_dimensionless.ncol_conga_ncol Tarski_neutral_dimensionless_axioms S3B angle_construction_1) obtain PP where S4A: "E Out P PP \ Cong E PP B X" by (metis G1 S4 os_distincts out_diff1 segment_construction_3) have S5: "P InAngle D E F" proof - have "PP InAngle DD E FF" proof - have Z3: "PP \ E" using S4A l6_3_1 by blast have Z4: "Bet DD PP FF" proof - have L1: "C B Q CongA P E F" proof - have K1: "B A OS C Q" using Col_perm P1 R1 S3B in_angle_one_side invert_one_side by blast have K2: "E D OS P F" by (simp add: S4 invert_one_side) have "C B A CongA P E D" by (simp add: S4 conga_comm) thus ?thesis using K1 K2 P1 l11_22b by auto qed have L2: "Cong DD FF A Q" proof - have "DD E FF CongA A B Q" proof - have L3: "A B Q CongA D E F" by (simp add: P1) then have L3A: "D E F CongA A B Q" using conga_sym by blast have L4: "E Out DD D" using G2 Out_cases by auto have L5: "E Out FF F" using G3G Out_cases by blast have L6: "B Out A A" using S3B not_col_distincts out_trivial by auto have "B Out Q Q" by (metis S3B not_col_distincts out_trivial) thus ?thesis using L3A L4 L5 L6 l11_10 by blast qed have L2B: "Cong DD E A B" using Cong_perm G2 by blast have "Cong E FF B Q" by (simp add: G3G) thus ?thesis using L2B \DD E FF CongA A B Q\ cong2_conga_cong by auto qed have L8: "Cong A X DD PP" proof - have L9: "A B X CongA DD E PP" proof - have L9B: "B Out A A" using S3B not_col_distincts out_trivial by blast have L9D: "E Out D D " using G3 out_trivial by auto have "E Out PP P" using Out_cases S4A by blast thus ?thesis using l11_10 S4 L9B G1 L9D using G2 Out_cases by blast qed have L10: "Cong A B DD E" using G2 not_cong_4321 by blast have "Cong B X E PP" using Cong_perm S4A by blast thus ?thesis using L10 L9 cong2_conga_cong by blast qed have "A X Q Cong3 DD PP FF" proof - have L12B: "Cong A Q DD FF" using L2 not_cong_3412 by blast have "Cong X Q PP FF" proof - have L13A: "X B Q CongA PP E FF" proof - have L13AC: "B Out Q Q" by (metis S3B col_trivial_2 out_trivial) have L13AD: "E Out PP P" by (simp add: S4A l6_6) have "E Out FF F" by (simp add: G3G l6_6) thus ?thesis using L1 G1 L13AC L13AD l11_10 by blast qed have L13B: "Cong X B PP E" using S4A not_cong_4321 by blast have "Cong B Q E FF" using G3G not_cong_3412 by blast thus ?thesis using L13A L13B cong2_conga_cong by auto qed thus ?thesis by (simp add: Cong3_def L12B L8) qed thus ?thesis using P2 l4_6 by blast qed have "PP = E \ E Out PP PP" using out_trivial by auto thus ?thesis using InAngle_def G3 G3B Z3 Z4 by auto qed thus ?thesis using G2 G3G S4A l11_25 by blast qed then have "A B C LeA D E F" using S4 LeA_def by blast } then have "A B C LeA D E F" using R2 \B Out A Q \ A B C LeA D E F\ \Bet A B Q \ A B C LeA D E F\ by blast } then have "A B C LeA D E F" using G4 \B Out A C \ A B C LeA D E F\ \Bet A B C \ A B C LeA D E F\ by blast } thus ?thesis using P2 \X = B \ A B C LeA D E F\ by blast qed lemma bet_in_angle_bet: assumes "Bet A B P" and "P InAngle A B C" shows "Bet A B C" by (metis (no_types) Col_def Col_perm assms(1) assms(2) col_in_angle_out not_bet_and_out) lemma lea_line: assumes "Bet A B P" and "A B P LeA A B C" shows "Bet A B C" by (metis Tarski_neutral_dimensionless.bet_conga__bet Tarski_neutral_dimensionless.l11_29_a Tarski_neutral_dimensionless_axioms assms(1) assms(2) bet_in_angle_bet) lemma eq_conga_out: assumes "A B A CongA D E F" shows "E Out D F" by (metis CongA_def assms l11_21_a out_trivial) lemma out_conga_out: assumes "B Out A C" and "A B C CongA D E F" shows "E Out D F" using assms(1) assms(2) l11_21_a by blast lemma conga_ex_cong3: assumes "A B C CongA A' B' C'" shows "\ AA CC. ((B Out A AA \ B Out C CC) \ AA B CC Cong3 A' B' C')" using out_diff2 by blast lemma conga_preserves_in_angle: assumes "A B C CongA A' B' C'" and "A B I CongA A' B' I'" and "I InAngle A B C" and "A' B' OS I' C'" shows "I' InAngle A' B' C'" proof - have P1: "A \ B" using assms(1) conga_diff1 by auto have P2: "B \ C" using assms(1) conga_diff2 by blast have P3: "A' \ B'" using assms(1) conga_diff45 by auto have P4: "B' \ C'" using assms(1) conga_diff56 by blast have P5: "I \ B" using assms(2) conga_diff2 by auto have P6: "I' \ B'" using assms(2) conga_diff56 by blast have P7: "Bet A B C \ B Out A C \ \ Col A B C" using l6_4_2 by blast { assume "Bet A B C" have Q1: "Bet A' B' C'" using \Bet A B C\ assms(1) assms(4) bet_col col124__nos col_conga_col by blast then have "I' InAngle A' B' C'" using assms(4) bet_col col124__nos by auto } { assume "B Out A C" then have "I' InAngle A' B' C'" by (metis P4 assms(2) assms(3) in_angle_out l11_21_a out321__inangle) } { assume Z1: "\ Col A B C" have Q2: "Bet A B I \ B Out A I \ \ Col A B I" by (simp add: or_bet_out) { assume "Bet A B I" then have "I' InAngle A' B' C'" using \Bet A B C \ I' InAngle A' B' C'\ assms(3) bet_in_angle_bet by blast } { assume "B Out A I" then have "I' InAngle A' B' C'" using P4 assms(2) l11_21_a out321__inangle by auto } { assume "\ Col A B I" obtain AA' where Q3: "B' Out A' AA' \ Cong B' AA' B A" using P1 P3 segment_construction_3 by presburger obtain CC' where Q4: "B' Out C' CC' \ Cong B' CC' B C" using P2 P4 segment_construction_3 by presburger obtain J where Q5: "Bet A J C \ (J = B \ B Out J I)" using InAngle_def assms(3) by auto have Q6: "B \ J" using Q5 Z1 bet_col by auto have Q7: "\ Col A B J" using Q5 Q6 \\ Col A B I\ col_permutation_2 col_transitivity_1 out_col by blast have "\ Col A' B' I'" by (metis assms(4) col123__nos) then have "\ C'. (A B J CongA A' B' C' \ A' B' OS C' I')" using Q7 angle_construction_1 by blast then obtain J' where Q8: "A B J CongA A' B' J' \ A' B' OS J' I'" by blast have Q9: "B' \ J'" using Q8 conga_diff56 by blast obtain JJ' where Q10: "B' Out J' JJ' \ Cong B' JJ' B J" using segment_construction_3 Q6 Q9 by blast have Q11: "\ Col A' B' J'" using Q8 col123__nos by blast have Q12: "A' \ JJ'" by (metis Col_perm Q10 Q11 out_col) have Q13: "B' \ JJ'" using Q10 out_distinct by blast have Q14: "\ Col A' B' JJ'" using Col_perm Q10 Q11 Q13 l6_16_1 out_col by blast have Q15: "A B C CongA AA' B' CC'" proof - have T2: "C \ B" using P2 by auto have T3: "AA' \ B'" using Out_def Q3 by blast have T4: "CC' \ B'" using Q4 out_distinct by blast have T5: "\ A' C' D' F'. (B Out A' A \ B Out C' C \ B' Out D' AA' \ B' Out F' CC' \Cong B A' B' D' \ Cong B C' B' F' \ Cong A' C' D' F')" by (smt Q3 Q4 Tarski_neutral_dimensionless.l11_4_1 Tarski_neutral_dimensionless_axioms assms(1) l6_6 l6_7) thus ?thesis using P1 T2 T3 T4 l11_4_2 by blast qed have Q16: "A' B' J' CongA A' B' JJ'" proof - have P9: "B' Out A' A'" by (simp add: P3 out_trivial) have "B' Out JJ' J'" using Out_cases Q10 by auto thus ?thesis using l11_10 by (simp add: P9 out2__conga) qed have Q17: "B' Out I' JJ' \ A' B' TS I' JJ'" proof - have "Coplanar A' I' B' J'" by (metis (full_types) Q8 ncoplanar_perm_3 os__coplanar) then have "Coplanar A' I' B' JJ'" using Q10 Q9 col_cop__cop out_col by blast then have R1: "Coplanar A' B' I' JJ'" using coplanar_perm_2 by blast have "A' B' I' CongA A' B' JJ'" proof - have R2: "A' B' I' CongA A B I" by (simp add: assms(2) conga_sym) have "A B I CongA A' B' JJ'" proof - have f1: "\p pa pb. \ p Out pa pb \ \ p Out pb pa \ p Out pa pb" using Out_cases by blast then have f2: "B' Out JJ' J'" using Q10 by blast have "B Out J I" by (metis Q5 Q6) thus ?thesis using f2 f1 by (meson P3 Q8 Tarski_neutral_dimensionless.l11_10 Tarski_neutral_dimensionless_axioms \\ Col A B I\ col_one_side_out col_trivial_2 one_side_reflexivity out_trivial) qed thus ?thesis using R2 conga_trans by blast qed thus ?thesis using R1 conga_cop__or_out_ts by blast qed { assume Z2: "B' Out I' JJ'" have Z3: "J B C CongA J' B' C'" proof - have R1: "B A OS J C" by (metis Q5 Q7 Z1 bet_out invert_one_side not_col_distincts out_one_side) have R2: "B' A' OS J' C'" by (meson Q10 Z2 assms(4) invert_one_side l6_6 one_side_symmetry out_out_one_side) have "J B A CongA J' B' A'" using Q8 conga_comm by blast thus ?thesis using assms(1) R1 R2 l11_22b by blast qed then have "I' InAngle A' B' C'" proof - have "A J C Cong3 AA' JJ' CC'" proof - have R8: "Cong A J AA' JJ'" proof - have R8A: "A B J CongA AA' B' JJ'" proof - have R8AB: "B Out A A" by (simp add: P1 out_trivial) have R8AC: "B Out J I" using Q5 Q6 by auto have R8AD: "B' Out AA' A'" using Out_cases Q3 by auto have "B' Out JJ' I'" using Out_cases Z2 by blast thus ?thesis using assms(2) R8AB R8AC R8AD l11_10 by blast qed have R8B: "Cong A B AA' B'" using Q3 not_cong_4321 by blast have R8C: "Cong B J B' JJ'" using Q10 not_cong_3412 by blast thus ?thesis using R8A R8B cong2_conga_cong by blast qed have LR8A: "Cong A C AA' CC'" using Q15 Q3 Q4 cong2_conga_cong cong_4321 cong_symmetry by blast have "Cong J C JJ' CC'" proof - have K1:"B' Out JJ' J'" using Out_cases Q10 by auto have "B' Out CC' C'" using Out_cases Q4 by auto then have "J' B' C' CongA JJ' B' CC'" using K1 by (simp add: out2__conga) then have LR9A: "J B C CongA JJ' B' CC'" using Z3 conga_trans by blast have LR9B: "Cong J B JJ' B'" using Q10 not_cong_4321 by blast have "Cong B C B' CC'" using Q4 not_cong_3412 by blast thus ?thesis using LR9A LR9B cong2_conga_cong by blast qed thus ?thesis using R8 LR8A by (simp add: Cong3_def) qed then have R10: "Bet AA' JJ' CC'" using Q5 l4_6 by blast have "JJ' InAngle AA' B' CC'" proof - have R11: "AA' \ B'" using Out_def Q3 by auto have R12: "CC' \ B'" using Out_def Q4 by blast have "Bet AA' JJ' CC' \ (JJ' = B' \ B' Out JJ' JJ')" using R10 out_trivial by auto thus ?thesis using InAngle_def Q13 R11 R12 by auto qed thus ?thesis using Z2 Q3 Q4 l11_25 by blast qed } { assume X1: "A' B' TS I' JJ'" have "A' B' OS I' J'" by (simp add: Q8 one_side_symmetry) then have X2: "B' A' OS I' JJ'" using Q10 invert_one_side out_out_one_side by blast then have "I' InAngle A' B' C'" using X1 invert_one_side l9_9 by blast } then have "I' InAngle A' B' C'" using Q17 \B' Out I' JJ' \ I' InAngle A' B' C'\ by blast } then have "I' InAngle A' B' C'" using Q2 \B Out A I \ I' InAngle A' B' C'\ \Bet A B I \ I' InAngle A' B' C'\ by blast } thus ?thesis using P7 \B Out A C \ I' InAngle A' B' C'\ \Bet A B C \ I' InAngle A' B' C'\ by blast qed lemma l11_30: assumes "A B C LeA D E F" and "A B C CongA A' B' C'" and "D E F CongA D' E' F'" shows "A' B' C' LeA D' E' F'" proof - obtain Q where P1: "C InAngle A B Q \ A B Q CongA D E F" using assms(1) l11_29_a by blast have P1A: "C InAngle A B Q" using P1 by simp have P1B: "A B Q CongA D E F" using P1 by simp have P2: "A \ B" using P1A inangle_distincts by auto have P3: "C \ B" using P1A inangle_distincts by blast have P4: "A' \ B'" using CongA_def assms(2) by blast have P5: "C' \ B'" using CongA_def assms(2) by auto have P6: "D \ E" using CongA_def P1B by blast have P7: "F \ E" using CongA_def P1B by blast have P8: "D' \ E'" using CongA_def assms(3) by blast have P9: "F' \ E'" using CongA_def assms(3) by blast have P10: "Bet A' B' C' \ B' Out A' C' \ \ Col A' B' C'" using or_bet_out by blast { assume "Bet A' B' C'" then have "\ Q'. (C' InAngle A' B' Q' \ A' B' Q' CongA D' E' F')" by (metis P1 P4 P5 P8 P9 assms(2) assms(3) bet_conga__bet bet_in_angle_bet conga_line conga_sym inangle3123) } { assume R1: "B' Out A' C'" obtain Q' where R2: "D' E' F' CongA A' B' Q'" using P4 P8 P9 angle_construction_3 by blast then have "C' InAngle A' B' Q'" using col_in_angle P1 R1 conga_diff56 out321__inangle by auto then have "\ Q'. (C' InAngle A' B' Q' \ A' B' Q' CongA D' E' F')" using R2 conga_sym by blast } { assume R3: "\ Col A' B' C'" have R3A: "Bet D' E' F' \ E' Out D' F' \ \ Col D' E' F'" using or_bet_out by blast { assume "Bet D' E' F'" have "\ Q'. (C' InAngle A' B' Q' \ A' B' Q' CongA D' E' F')" by (metis P4 P5 P8 P9 \Bet D' E' F'\ conga_line in_angle_line point_construction_different) } { assume R4A: "E' Out D' F'" obtain Q' where R4: "D' E' F' CongA A' B' Q'" using P4 P8 P9 angle_construction_3 by blast then have R5: "B' Out A' Q'" using out_conga_out R4A by blast have R6: "A B Q CongA D' E' F'" using P1 assms(3) conga_trans by blast then have R7: "B Out A Q" using out_conga_out R4A R4 using conga_sym by blast have R8: "B Out A C" using P1A R7 in_angle_out by blast then have R9: "B' Out A' C'" using out_conga_out assms(2) by blast have "\ Q'. (C' InAngle A' B' Q' \ A' B' Q' CongA D' E' F')" by (simp add: R9 \B' Out A' C' \ \Q'. C' InAngle A' B' Q' \ A' B' Q' CongA D' E' F'\) } { assume "\ Col D' E' F'" obtain QQ where S1: "D' E' F' CongA A' B' QQ \ A' B' OS QQ C'" using R3 \\ Col D' E' F'\ angle_construction_1 by blast have S1A: "A B Q CongA A' B' QQ" using S1 using P1 assms(3) conga_trans by blast have "A' B' OS C' QQ" using S1 by (simp add: S1 one_side_symmetry) then have S2: "C' InAngle A' B' QQ" using conga_preserves_in_angle S1A using P1A assms(2) by blast have S3: "A' B' QQ CongA D' E' F'" by (simp add: S1 conga_sym) then have "\ Q'. (C' InAngle A' B' Q' \ A' B' Q' CongA D' E' F')" using S2 by auto } then have "\ Q'. (C' InAngle A' B' Q' \ A' B' Q' CongA D' E' F')" using R3A \E' Out D' F' \ \Q'. C' InAngle A' B' Q' \ A' B' Q' CongA D' E' F'\ \Bet D' E' F' \ \Q'. C' InAngle A' B' Q' \ A' B' Q' CongA D' E' F'\ by blast } thus ?thesis using l11_29_b using P10 \B' Out A' C' \ \Q'. C' InAngle A' B' Q' \ A' B' Q' CongA D' E' F'\ \Bet A' B' C' \ \Q'. C' InAngle A' B' Q' \ A' B' Q' CongA D' E' F'\ by blast qed lemma l11_31_1: assumes "B Out A C" and "D \ E" and "F \ E" shows "A B C LeA D E F" by (metis (full_types) LeA_def assms(1) assms(2) assms(3) l11_21_b out321__inangle segment_construction_3) lemma l11_31_2: assumes "A \ B" and "C \ B" and "D \ E" and "F \ E" and "Bet D E F" shows "A B C LeA D E F" by (metis LeA_def angle_construction_3 assms(1) assms(2) assms(3) assms(4) assms(5) conga_diff56 in_angle_line) lemma lea_refl: assumes "A \ B" and "C \ B" shows "A B C LeA A B C" by (meson assms(1) assms(2) conga_refl l11_29_b out341__inangle out_trivial) lemma conga__lea: assumes "A B C CongA D E F" shows "A B C LeA D E F" by (metis Tarski_neutral_dimensionless.conga_diff1 Tarski_neutral_dimensionless.conga_diff2 Tarski_neutral_dimensionless.l11_30 Tarski_neutral_dimensionless_axioms assms conga_refl lea_refl) lemma conga__lea456123: assumes "A B C CongA D E F" shows "D E F LeA A B C" by (simp add: Tarski_neutral_dimensionless.conga__lea Tarski_neutral_dimensionless_axioms assms conga_sym) lemma lea_left_comm: assumes "A B C LeA D E F" shows "C B A LeA D E F" by (metis assms conga_pseudo_refl conga_refl l11_30 lea_distincts) lemma lea_right_comm: assumes "A B C LeA D E F" shows "A B C LeA F E D" by (meson assms conga_right_comm l11_29_a l11_29_b) lemma lea_comm: assumes"A B C LeA D E F" shows "C B A LeA F E D" using assms lea_left_comm lea_right_comm by blast lemma lta_left_comm: assumes "A B C LtA D E F" shows "C B A LtA D E F" by (meson LtA_def Tarski_neutral_dimensionless.conga_left_comm Tarski_neutral_dimensionless.lea_left_comm Tarski_neutral_dimensionless_axioms assms) lemma lta_right_comm: assumes "A B C LtA D E F" shows "A B C LtA F E D" by (meson Tarski_neutral_dimensionless.LtA_def Tarski_neutral_dimensionless.conga_comm Tarski_neutral_dimensionless.lea_comm Tarski_neutral_dimensionless.lta_left_comm Tarski_neutral_dimensionless_axioms assms) lemma lta_comm: assumes "A B C LtA D E F" shows "C B A LtA F E D" using assms lta_left_comm lta_right_comm by blast lemma lea_out4__lea: assumes "A B C LeA D E F" and "B Out A A'" and "B Out C C'" and "E Out D D'" and "E Out F F'" shows "A' B C' LeA D' E F'" using assms(1) assms(2) assms(3) assms(4) assms(5) l11_30 l6_6 out2__conga by auto lemma lea121345: assumes "A \ B" and "C \ D" and "D \ E" shows "A B A LeA C D E" using assms(1) assms(2) assms(3) l11_31_1 out_trivial by auto lemma inangle__lea: assumes "P InAngle A B C" shows "A B P LeA A B C" by (metis Tarski_neutral_dimensionless.l11_29_b Tarski_neutral_dimensionless_axioms assms conga_refl inangle_distincts) lemma inangle__lea_1: assumes "P InAngle A B C" shows "P B C LeA A B C" by (simp add: Tarski_neutral_dimensionless.inangle__lea Tarski_neutral_dimensionless.lea_comm Tarski_neutral_dimensionless_axioms assms l11_24) lemma inangle__lta: assumes "\ Col P B C" and "P InAngle A B C" shows "A B P LtA A B C" by (metis LtA_def TS_def Tarski_neutral_dimensionless.conga_cop__or_out_ts Tarski_neutral_dimensionless.conga_os__out Tarski_neutral_dimensionless.inangle__lea Tarski_neutral_dimensionless.ncol_conga_ncol Tarski_neutral_dimensionless_axioms assms(1) assms(2) col_one_side_out col_trivial_3 in_angle_one_side inangle__coplanar invert_two_sides l11_21_b ncoplanar_perm_12 not_col_permutation_3 one_side_reflexivity) lemma in_angle_trans: assumes "C InAngle A B D" and "D InAngle A B E" shows "C InAngle A B E" proof - obtain CC where P1: "Bet A CC D \ (CC = B \ B Out CC C)" using InAngle_def assms(1) by auto obtain DD where P2: "Bet A DD E \ (DD = B \ B Out DD D)" using InAngle_def assms(2) by auto then have P3: "Bet A DD E" by simp have P4: "DD = B \ B Out DD D" using P2 by simp { assume "CC = B \ DD = B" then have "C InAngle A B E" using InAngle_def P2 assms(1) assms(2) by auto } { assume "CC = B \ B Out DD D" then have "C InAngle A B E" by (metis InAngle_def P1 assms(1) assms(2) bet_in_angle_bet) } { assume "B Out CC C \ DD = B" then have "C InAngle A B E" by (metis Out_def P2 assms(2) in_angle_line inangle_distincts) } { assume P3: "B Out CC C \ B Out DD D" then have P3A: "B Out CC C" by simp have P3B: "B Out DD D" using P3 by simp have "C InAngle A B DD" using P3 assms(1) inangle_distincts l11_25 out_trivial by blast then obtain CC' where T1: "Bet A CC' DD \ (CC' = B \ B Out CC' C)" using InAngle_def by auto { assume "CC' = B" then have "C InAngle A B E" by (metis P2 P3 T1 assms(2) between_exchange4 in_angle_line inangle_distincts out_diff2) } { assume "B Out CC' C" then have "C InAngle A B E" by (metis InAngle_def P2 T1 assms(1) assms(2) between_exchange4) } then have "C InAngle A B E" using T1 \CC' = B \ C InAngle A B E\ by blast } thus ?thesis using P1 P2 \B Out CC C \ DD = B \ C InAngle A B E\ \CC = B \ B Out DD D \ C InAngle A B E\ \CC = B \ DD = B \ C InAngle A B E\ by blast qed lemma lea_trans: assumes "A B C LeA A1 B1 C1" and "A1 B1 C1 LeA A2 B2 C2" shows "A B C LeA A2 B2 C2" proof - obtain P1 where T1: "P1 InAngle A1 B1 C1 \ A B C CongA A1 B1 P1" using LeA_def assms(1) by auto obtain P2 where T2: "P2 InAngle A2 B2 C2 \ A1 B1 C1 CongA A2 B2 P2" using LeA_def assms(2) by blast have T3: "A \ B" using CongA_def T1 by auto have T4: "C \ B" using CongA_def T1 by blast have T5: "A1 \ B1" using T1 inangle_distincts by blast have T6: "C1 \ B1" using T1 inangle_distincts by blast have T7: "A2 \ B2" using T2 inangle_distincts by blast have T8: "C2 \ B2" using T2 inangle_distincts by blast have T9: "Bet A B C \ B Out A C \ \ Col A B C" using not_out_bet by auto { assume "Bet A B C" then have "A B C LeA A2 B2 C2" by (metis T1 T2 T3 T4 T7 T8 bet_conga__bet bet_in_angle_bet l11_31_2) } { assume "B Out A C" then have "A B C LeA A2 B2 C2" by (simp add: T7 T8 l11_31_1) } { assume H1: "\ Col A B C" have T10: "Bet A2 B2 C2 \ B2 Out A2 C2 \ \ Col A2 B2 C2" using not_out_bet by auto { assume "Bet A2 B2 C2" then have "A B C LeA A2 B2 C2" by (simp add: T3 T4 T7 T8 l11_31_2) } { assume T10A: "B2 Out A2 C2" have "B Out A C" proof - have "B1 Out A1 P1" proof - have "B1 Out A1 C1" using T2 conga_sym T2 T10A in_angle_out out_conga_out by blast thus ?thesis using T1 in_angle_out by blast qed thus ?thesis using T1 conga_sym l11_21_a by blast qed then have "A B C LeA A2 B2 C2" using \B Out A C \ A B C LeA A2 B2 C2\ by blast } { assume T12: "\ Col A2 B2 C2" obtain P where T13: "A B C CongA A2 B2 P \ A2 B2 OS P C2" using T12 H1 angle_construction_1 by blast have T14: "A2 B2 OS P2 C2" proof - have "\ Col B2 A2 P2" proof - have "B2 \ A2" using T7 by auto { assume H2: "P2 = A2" have "A2 B2 A2 CongA A1 B1 C1" using T2 H2 conga_sym by blast then have "B1 Out A1 C1" using eq_conga_out by blast then have "B1 Out A1 P1" using T1 in_angle_out by blast then have "B Out A C" using T1 conga_sym out_conga_out by blast then have False using Col_cases H1 out_col by blast } then have "P2 \ A2" by blast have "Bet A2 B2 P2 \ B2 Out A2 P2 \ \ Col A2 B2 P2" using not_out_bet by auto { assume H4: "Bet A2 B2 P2" then have "Bet A2 B2 C2" using T2 bet_in_angle_bet by blast then have "Col B2 A2 P2 \ False" using Col_def T12 by blast then have "\ Col B2 A2 P2" using H4 bet_col not_col_permutation_4 by blast } { assume H5: "B2 Out A2 P2" then have "B1 Out A1 C1" using T2 conga_sym out_conga_out by blast then have "B1 Out A1 P1" using T1 in_angle_out by blast then have "B Out A C" using H1 T1 ncol_conga_ncol not_col_permutation_4 out_col by blast then have "\ Col B2 A2 P2" using Col_perm H1 out_col by blast } { assume "\ Col A2 B2 P2" then have "\ Col B2 A2 P2" using Col_perm by blast } thus ?thesis using \B2 Out A2 P2 \ \ Col B2 A2 P2\ \Bet A2 B2 P2 \ \ Col B2 A2 P2\ \Bet A2 B2 P2 \ B2 Out A2 P2 \ \ Col A2 B2 P2\ by blast qed thus ?thesis by (simp add: T2 T12 in_angle_one_side) qed have S1: "A2 B2 OS P P2" using T13 T14 one_side_symmetry one_side_transitivity by blast have "A1 B1 P1 CongA A2 B2 P" using conga_trans conga_sym T1 T13 by blast then have "P InAngle A2 B2 P2" using conga_preserves_in_angle T2 T1 S1 by blast then have "P InAngle A2 B2 C2" using T2 in_angle_trans by blast then have "A B C LeA A2 B2 C2" using T13 LeA_def by blast } then have "A B C LeA A2 B2 C2" using T10 \B2 Out A2 C2 \ A B C LeA A2 B2 C2\ \Bet A2 B2 C2 \ A B C LeA A2 B2 C2\ by blast } thus ?thesis using T9 \B Out A C \ A B C LeA A2 B2 C2\ \Bet A B C \ A B C LeA A2 B2 C2\ by blast qed lemma in_angle_asym: assumes "D InAngle A B C" and "C InAngle A B D" shows "A B C CongA A B D" proof - obtain CC where P1: "Bet A CC D \ (CC = B \ B Out CC C)" using InAngle_def assms(2) by auto obtain DD where P2: "Bet A DD C \ (DD = B \ B Out DD D)" using InAngle_def assms(1) by auto { assume "(CC = B) \ (DD = B)" then have "A B C CongA A B D" by (metis P1 P2 assms(2) conga_line inangle_distincts) } { assume "(CC = B) \ (B Out DD D)" then have "A B C CongA A B D" by (metis P1 assms(1) bet_in_angle_bet conga_line inangle_distincts) } { assume "(B Out CC C) \ (DD = B)" then have "A B C CongA A B D" by (metis P2 assms(2) bet_in_angle_bet conga_line inangle_distincts) } { assume V1: "(B Out CC C) \ (B Out DD D)" obtain X where P3: "Bet CC X C \ Bet DD X D" using P1 P2 between_symmetry inner_pasch by blast then have "B Out X D" using V1 out_bet_out_2 by blast then have "B Out C D" using P3 V1 out2_bet_out by blast then have "A B C CongA A B D" using assms(2) inangle_distincts l6_6 out2__conga out_trivial by blast } thus ?thesis using P1 P2 using \B Out CC C \ DD = B \ A B C CongA A B D\ \CC = B \ B Out DD D \ A B C CongA A B D\ \CC = B \ DD = B \ A B C CongA A B D\ by blast qed lemma lea_asym: assumes "A B C LeA D E F" and "D E F LeA A B C" shows "A B C CongA D E F" proof cases assume P1: "Col A B C" { assume P1A: "Bet A B C" have P2: "D \ E" using assms(1) lea_distincts by blast have P3: "F \ E" using assms(2) lea_distincts by auto have P4: "A \ B" using assms(1) lea_distincts by auto have P5: "C \ B" using assms(2) lea_distincts by blast obtain P where P6: "P InAngle D E F \ A B C CongA D E P" using LeA_def assms(1) by blast then have "A B C CongA D E P" by simp then have "Bet D E P" using P1 P1A bet_conga__bet by blast then have "Bet D E F" using P6 bet_in_angle_bet by blast then have "A B C CongA D E F" by (metis Tarski_neutral_dimensionless.bet_conga__bet Tarski_neutral_dimensionless.conga_line Tarski_neutral_dimensionless.l11_29_a Tarski_neutral_dimensionless_axioms P2 P3 P4 P5 assms(2) bet_in_angle_bet) } { assume T1: "\ Bet A B C" then have T2: "B Out A C" using P1 not_out_bet by auto obtain P where T3: "P InAngle A B C \ D E F CongA A B P" using LeA_def assms(2) by blast then have T3A: "P InAngle A B C" by simp have T3B: "D E F CongA A B P" using T3 by simp have T4: "E Out D F" proof - have T4A: "B Out A P" using T2 T3 in_angle_out by blast have "A B P CongA D E F" by (simp add: T3 conga_sym) thus ?thesis using T4A l11_21_a by blast qed then have "A B C CongA D E F" by (simp add: T2 l11_21_b) } thus ?thesis using \Bet A B C \ A B C CongA D E F\ by blast next assume T5: "\ Col A B C" obtain Q where T6: "C InAngle A B Q \ A B Q CongA D E F" using assms(1) l11_29_a by blast then have T6A: "C InAngle A B Q" by simp have T6B: "A B Q CongA D E F" by (simp add: T6) obtain P where T7: "P InAngle A B C \ D E F CongA A B P" using LeA_def assms(2) by blast then have T7A: "P InAngle A B C" by simp have T7B: "D E F CongA A B P" by (simp add: T7) have T13: "A B Q CongA A B P" using T6 T7 conga_trans by blast have T14: "Bet A B Q \ B Out A Q \ \ Col A B Q" using not_out_bet by auto { assume R1: "Bet A B Q" then have "A B C CongA D E F" using T13 T5 T7 bet_col bet_conga__bet bet_in_angle_bet by blast } { assume R2: "B Out A Q" then have "A B C CongA D E F" using T6 in_angle_out l11_21_a l11_21_b by blast } { assume R3: "\ Col A B Q" have R3A: "Bet A B P \ B Out A P \ \ Col A B P" using not_out_bet by blast { assume R3AA: "Bet A B P" then have "A B C CongA D E F" using T5 T7 bet_col bet_in_angle_bet by blast } { assume R3AB: "B Out A P" then have "A B C CongA D E F" by (meson Col_cases R3 T13 ncol_conga_ncol out_col) } { assume R3AC: "\ Col A B P" have R3AD: "B Out P Q \ A B TS P Q" proof - have "Coplanar A B P Q" using T6A T7A coplanar_perm_8 in_angle_trans inangle__coplanar by blast thus ?thesis by (simp add: T13 conga_sym conga_cop__or_out_ts) qed { assume "B Out P Q" then have "C InAngle A B P" by (meson R3 T6A bet_col between_symmetry l11_24 l11_25_aux) then have "A B C CongA A B P" by (simp add: T7A in_angle_asym) then have "A B C CongA D E F" by (meson T7B Tarski_neutral_dimensionless.conga_sym Tarski_neutral_dimensionless.conga_trans Tarski_neutral_dimensionless_axioms) } { assume W1: "A B TS P Q" have "A B OS P Q" using Col_perm R3 R3AC T6A T7A in_angle_one_side in_angle_trans by blast then have "A B C CongA D E F" using W1 l9_9 by blast } then have "A B C CongA D E F" using R3AD \B Out P Q \ A B C CongA D E F\ by blast } then have "A B C CongA D E F" using R3A \B Out A P \ A B C CongA D E F\ \Bet A B P \ A B C CongA D E F\ by blast } thus ?thesis using T14 \B Out A Q \ A B C CongA D E F\ \Bet A B Q \ A B C CongA D E F\ by blast qed lemma col_lta__bet: assumes "Col X Y Z" and "A B C LtA X Y Z" shows "Bet X Y Z" proof - have "A B C LeA X Y Z \ \ A B C CongA X Y Z" using LtA_def assms(2) by auto then have "Y Out X Z \ False" using Tarski_neutral_dimensionless.lea_asym Tarski_neutral_dimensionless.lea_distincts Tarski_neutral_dimensionless_axioms l11_31_1 by fastforce thus ?thesis using not_out_bet assms(1) by blast qed lemma col_lta__out: assumes "Col A B C" and "A B C LtA X Y Z" shows "B Out A C" proof - have "A B C LeA X Y Z \ \ A B C CongA X Y Z" using LtA_def assms(2) by auto thus ?thesis by (metis assms(1) l11_31_2 lea_asym lea_distincts or_bet_out) qed lemma lta_distincts: assumes "A B C LtA D E F" shows "A\B \ C\B \ D\E \ F\E \ D \ F" by (metis LtA_def assms bet_neq12__neq col_lta__bet lea_distincts not_col_distincts) lemma gta_distincts: assumes "A B C GtA D E F" shows "A\B \ C\B \ D\E \ F\E \ A \ C" using GtA_def assms lta_distincts by presburger lemma acute_distincts: assumes "Acute A B C" shows "A\B \ C\B" using Acute_def assms lta_distincts by blast lemma obtuse_distincts: assumes "Obtuse A B C" shows "A\B \ C\B \ A \ C" using Obtuse_def assms lta_distincts by blast lemma two_sides_in_angle: assumes "B \ P'" and "B P TS A C" and "Bet P B P'" shows "P InAngle A B C \ P' InAngle A B C" proof - obtain T where P1: "Col T B P \ Bet A T C" using TS_def assms(2) by auto have P2: "A \ B" using assms(2) ts_distincts by blast have P3: "C \ B" using assms(2) ts_distincts by blast show ?thesis proof cases assume "B = T" thus ?thesis using P1 P2 P3 assms(1) in_angle_line by auto next assume "B \ T" thus ?thesis by (metis InAngle_def P1 assms(1) assms(2) assms(3) between_symmetry l6_3_2 or_bet_out ts_distincts) qed qed lemma in_angle_reverse: assumes "A' \ B" and "Bet A B A'" and "C InAngle A B D" shows "D InAngle A' B C" proof - have P1: "A \ B" using assms(3) inangle_distincts by auto have P2: "D \ B" using assms(3) inangle_distincts by blast have P3: "C \ B" using assms(3) inangle_distincts by auto show ?thesis proof cases assume "Col B A C" thus ?thesis by (smt P1 P2 P3 assms(1) assms(2) assms(3) bet_in_angle_bet between_inner_transitivity between_symmetry in_angle_line l6_3_2 out321__inangle outer_transitivity_between third_point) next assume P4: "\ Col B A C" thus ?thesis proof cases assume "Col B D C" thus ?thesis by (smt P2 P4 assms(1) assms(2) assms(3) bet_col1 col2__eq col_permutation_2 in_angle_one_side l9_19_R1 out341__inangle) next assume P5: "\ Col B D C" have P6: "C B TS A D" using P4 P5 assms(3) in_angle_two_sides by auto obtain X where P7: "Bet A X D \ (X = B \ B Out X C)" using InAngle_def assms(3) by auto have P8: "X = B \ D InAngle A' B C" using Out_def P1 P2 P3 P7 assms(1) assms(2) l5_2 out321__inangle by auto { assume P9: "B Out X C" have P10: "C \ B" by (simp add: P3) have P10A: "\ Col B A C" by (simp add: P4) have P10B: "\ Col B D C" by (simp add: P5) have P10C: "C InAngle D B A" by (simp add: assms(3) l11_24) { assume "Col D B A" have "Col B A C" proof - have "B \ X" using P9 out_distinct by blast have "Col B X A" by (meson Bet_perm P10C P5 P7 \Col D B A\ bet_col1 col_permutation_3 in_angle_out or_bet_out out_col) have "Col B X C" by (simp add: P9 out_col) thus ?thesis using \B \ X\ \Col B X A\ col_transitivity_1 by blast qed then have False by (simp add: P4) } then have P10E: "\ Col D B A" by auto have P11: "D B OS C A" by (simp add: P10C P10E P5 in_angle_one_side) have P12: "\ Col A D B" using Col_cases P10E by auto have P13: "\ Col A' D B" by (metis Col_def \Col D B A \ False\ assms(1) assms(2) col_transitivity_1) have P14: "D B TS A A'" using P12 P13 TS_def assms(2) col_trivial_3 by blast have P15: "D B TS C A'" using P11 P14 l9_8_2 one_side_symmetry by blast have P16: "\ Col C D B" by (simp add: P5 not_col_permutation_3) obtain Y where P17: "Col Y D B \ Bet C Y A'" using P15 TS_def by auto have P18: "Bet A' Y C" using Bet_perm P17 by blast { assume S1: "Y \ B" have S2: "Col D B Y" using P17 not_col_permutation_2 by blast then have S3: "Bet D B Y \ Bet B Y D \ Bet Y D B" using Col_def S2 by auto { assume S4: "Bet D B Y" have S5: "C B OS A' Y" by (metis P17 P18 P5 S1 bet_out_1 col_transitivity_2 l6_6 not_col_permutation_3 not_col_permutation_5 out_one_side) have S6: "C B TS Y D" by (metis Bet_perm P16 P17 S1 S4 bet__ts col3 col_trivial_3 invert_two_sides not_col_permutation_1) have "C B TS A A'" by (metis (full_types) P4 assms(1) assms(2) bet__ts invert_two_sides not_col_permutation_5) then have "C B TS Y A" using S5 l9_2 l9_8_2 by blast then have S9: "C B OS A D" using P6 S6 l9_8_1 l9_9 by blast then have "B Out Y D" using P6 S9 l9_9 by auto } { assume "Bet B Y D" then have "B Out Y D" by (simp add: S1 bet_out) } { assume "Bet Y D B" then have "B Out Y D" by (simp add: P2 bet_out_1 l6_6) } have "B Out Y D" using S3 \Bet B Y D \ B Out Y D\ \Bet D B Y \ B Out Y D\ \Bet Y D B \ B Out Y D\ by blast } then have P19: "(Y = B \ B Out Y D)" by auto have "D InAngle A' B C" using InAngle_def P18 P19 P2 P3 assms(1) by auto } thus ?thesis using P7 P8 by blast qed qed qed lemma in_angle_trans2: assumes "C InAngle A B D" and "D InAngle A B E" shows "D InAngle C B E" proof - obtain pp :: "'p \ 'p \ 'p" where f1: "\p pa. Bet p pa (pp p pa) \ pa \ (pp p pa)" using point_construction_different by moura then have f2: "\p. C InAngle D B (pp p B) \ \ D InAngle p B A" by (metis assms(1) in_angle_reverse in_angle_trans l11_24) have f3: "D InAngle E B A" using assms(2) l11_24 by blast then have "E \ B" by (simp add: inangle_distincts) thus ?thesis using f3 f2 f1 by (meson Bet_perm in_angle_reverse l11_24) qed lemma l11_36_aux1: assumes "A \ B" and "A' \ B" and "D \ E" and "D' \ E" and "Bet A B A'" and "Bet D E D'" and "A B C LeA D E F" shows "D' E F LeA A' B C" proof - obtain P where P1: "C InAngle A B P \ A B P CongA D E F" using assms(7) l11_29_a by blast thus ?thesis by (metis LeA_def Tarski_neutral_dimensionless.l11_13 Tarski_neutral_dimensionless_axioms assms(2) assms(4) assms(5) assms(6) conga_sym in_angle_reverse) qed lemma l11_36_aux2: assumes "A \ B" and "A' \ B" and "D \ E" and "D' \ E" and "Bet A B A'" and "Bet D E D'" and "D' E F LeA A' B C" shows "A B C LeA D E F" by (metis Bet_cases assms(1) assms(3) assms(5) assms(6) assms(7) l11_36_aux1 lea_distincts) lemma l11_36: assumes "A \ B" and "A' \ B" and "D \ E" and "D' \ E" and "Bet A B A'" and "Bet D E D'" shows "A B C LeA D E F \ D' E F LeA A' B C" using assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) l11_36_aux1 l11_36_aux2 by auto lemma l11_41_aux: assumes "\ Col A B C" and "Bet B A D" and "A \ D" shows "A C B LtA C A D" proof - obtain M where P1: "M Midpoint A C" using midpoint_existence by auto obtain P where P2: "M Midpoint B P" using symmetric_point_construction by auto have P3: "A C B Cong3 C A P" by (smt Cong3_def P1 P2 assms(1) l7_13_R1 l7_2 midpoint_distinct_1 not_col_distincts) have P4: "A \ C" using assms(1) col_trivial_3 by blast have P5: "B \ C" using assms(1) col_trivial_2 by blast have P7: "A \ M" using P1 P4 is_midpoint_id by blast have P8: "A C B CongA C A P" by (simp add: P3 P4 P5 cong3_conga) have P8A: "Bet D A B" using Bet_perm assms(2) by blast have P8B: "Bet P M B" by (simp add: P2 between_symmetry midpoint_bet) then obtain X where P9: "Bet A X P \ Bet M X D" using P8A inner_pasch by blast have P9A: "Bet A X P" by (simp add: P9) have P9B: "Bet M X D" by (simp add: P9) have P10A: "P InAngle C A D" proof - have K1: "P InAngle M A D" by (metis InAngle_def P3 P5 P7 P9 assms(3) bet_out cong3_diff2) have K2: "A Out C M" using Out_def P1 P4 P7 midpoint_bet by auto have K3: "A Out D D" using assms(3) out_trivial by auto have "A Out P P" using K1 inangle_distincts out_trivial by auto thus ?thesis using K1 K2 K3 l11_25 by blast qed then have P10: "A C B LeA C A D" using LeA_def P8 by auto { assume K5: "A C B CongA C A D" then have K6: "C A D CongA C A P" using P8 conga_sym conga_trans by blast have K7: "Coplanar C A D P" using P10A inangle__coplanar ncoplanar_perm_18 by blast then have K8: "A Out D P \ C A TS D P" by (simp add: K6 conga_cop__or_out_ts) { assume "A Out D P" then have "Col M B A" by (meson P8A P8B bet_col1 bet_out__bet between_symmetry not_col_permutation_4) then have K8F: "Col A M B" using not_col_permutation_1 by blast have "Col A M C" by (simp add: P1 bet_col midpoint_bet) then have "False" using K8F P7 assms(1) col_transitivity_1 by blast } then have K9: "\ A Out D P" by auto { assume V1: "C A TS D P" then have V3: "A C TS B P" by (metis P10A P8A assms(1) col_trivial_1 col_trivial_2 in_angle_reverse in_angle_two_sides invert_two_sides l11_24 l9_18 not_col_permutation_5) have "A C TS B D" by (simp add: assms(1) assms(2) assms(3) bet__ts not_col_permutation_5) then have "A C OS D P" using V1 V3 invert_two_sides l9_8_1 l9_9 by blast then have "False" using V1 invert_one_side l9_9 by blast } then have "\ C A TS D P" by auto then have "False" using K8 K9 by auto } then have "\ A C B CongA C A D" by auto thus ?thesis by (simp add: LtA_def P10) qed lemma l11_41: assumes "\ Col A B C" and "Bet B A D" and "A \ D" shows "A C B LtA C A D \ A B C LtA C A D" proof - have P1: "A C B LtA C A D" using assms(1) assms(2) assms(3) l11_41_aux by auto have "A B C LtA C A D" proof - obtain E where T1: "Bet C A E \ Cong A E C A" using segment_construction by blast have T1A: "Bet C A E" using T1 by simp have T1B: "Cong A E C A" using T1 by simp have T2: "A B C LtA B A E" using T1 assms(1) cong_reverse_identity l11_41_aux not_col_distincts not_col_permutation_5 by blast have T3: "B A C CongA C A B" by (metis assms(1) conga_pseudo_refl not_col_distincts) have T3A: "D A C CongA E A B" by (metis CongA_def T1 T3 assms(2) assms(3) cong_reverse_identity l11_13) then have T4: "B A E CongA C A D" using conga_comm conga_sym by blast have "A B C CongA A B C" using T2 Tarski_neutral_dimensionless.conga_refl Tarski_neutral_dimensionless.lta_distincts Tarski_neutral_dimensionless_axioms by fastforce then have T5: "A B C LeA C A D" by (meson T2 T4 Tarski_neutral_dimensionless.LtA_def Tarski_neutral_dimensionless.l11_30 Tarski_neutral_dimensionless_axioms) have "\ A B C CongA C A D" by (meson T2 Tarski_neutral_dimensionless.LtA_def Tarski_neutral_dimensionless.conga_right_comm Tarski_neutral_dimensionless.conga_trans Tarski_neutral_dimensionless_axioms T3A) thus ?thesis by (simp add: LtA_def T5) qed thus ?thesis by (simp add: P1) qed lemma not_conga: assumes "A B C CongA A' B' C'" and "\ A B C CongA D E F" shows "\ A' B' C' CongA D E F" by (meson assms(1) assms(2) conga_trans) lemma not_conga_sym: assumes "\ A B C CongA D E F" shows "\ D E F CongA A B C" using assms conga_sym by blast lemma not_and_lta: shows "\ (A B C LtA D E F \ D E F LtA A B C)" proof - { assume P1: "A B C LtA D E F \ D E F LtA A B C" then have "A B C CongA D E F" using LtA_def lea_asym by blast then have "False" using LtA_def P1 by blast } thus ?thesis by auto qed lemma conga_preserves_lta: assumes "A B C CongA A' B' C'" and "D E F CongA D' E' F'" and "A B C LtA D E F" shows "A' B' C' LtA D' E' F'" by (meson Tarski_neutral_dimensionless.LtA_def Tarski_neutral_dimensionless.conga_trans Tarski_neutral_dimensionless.l11_30 Tarski_neutral_dimensionless.not_conga_sym Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3)) lemma lta_trans: assumes "A B C LtA A1 B1 C1" and "A1 B1 C1 LtA A2 B2 C2" shows "A B C LtA A2 B2 C2" proof - have P1: "A B C LeA A2 B2 C2" by (meson LtA_def assms(1) assms(2) lea_trans) { assume "A B C CongA A2 B2 C2" then have "False" by (meson Tarski_neutral_dimensionless.LtA_def Tarski_neutral_dimensionless.lea_asym Tarski_neutral_dimensionless.lea_trans Tarski_neutral_dimensionless_axioms assms(1) assms(2) conga__lea456123) } thus ?thesis using LtA_def P1 by blast qed lemma obtuse_sym: assumes "Obtuse A B C" shows "Obtuse C B A" by (meson Obtuse_def Tarski_neutral_dimensionless.lta_right_comm Tarski_neutral_dimensionless_axioms assms) lemma acute_sym: assumes "Acute A B C" shows "Acute C B A" by (meson Acute_def Tarski_neutral_dimensionless.lta_left_comm Tarski_neutral_dimensionless_axioms assms) lemma acute_col__out: assumes "Col A B C" and "Acute A B C" shows "B Out A C" by (meson Tarski_neutral_dimensionless.Acute_def Tarski_neutral_dimensionless_axioms assms(1) assms(2) col_lta__out) lemma col_obtuse__bet: assumes "Col A B C" and "Obtuse A B C" shows "Bet A B C" using Obtuse_def assms(1) assms(2) col_lta__bet by blast lemma out__acute: assumes "B Out A C" shows "Acute A B C" proof - have P1: "A \ B" using assms out_diff1 by auto then obtain D where P3: "B D Perp A B" using perp_exists by blast then have P4: "B \ D" using perp_distinct by auto have P5: "Per A B D" by (simp add: P3 l8_2 perp_per_1) have P6: "A B C LeA A B D" using P1 P4 assms l11_31_1 by auto { assume "A B C CongA A B D" then have "False" by (metis Col_cases P1 P4 P5 assms col_conga_col l8_9 out_col) } then have "A B C LtA A B D" using LtA_def P6 by auto thus ?thesis using P5 Acute_def by auto qed lemma bet__obtuse: assumes "Bet A B C" and "A \ B" and "B \ C" shows "Obtuse A B C" proof - obtain D where P1: "B D Perp A B" using assms(2) perp_exists by blast have P5: "B \ D" using P1 perp_not_eq_1 by auto have P6: "Per A B D" using P1 Perp_cases perp_per_1 by blast have P7: "A B D LeA A B C" using assms(2) assms(3) P5 assms(1) l11_31_2 by auto { assume "A B D CongA A B C" then have "False" using assms(2) P5 P6 assms(1) bet_col ncol_conga_ncol per_not_col by blast } then have "A B D LtA A B C" using LtA_def P7 by blast thus ?thesis using Obtuse_def P6 by blast qed lemma l11_43_aux: assumes "A \ B" and "A \ C" and "Per B A C \ Obtuse B A C" shows "Acute A B C" proof cases assume P1: "Col A B C" { assume "Per B A C" then have "Acute A B C" using Col_cases P1 assms(1) assms(2) per_col_eq by blast } { assume "Obtuse B A C" then have "Bet B A C" using P1 col_obtuse__bet col_permutation_4 by blast then have "Acute A B C" by (simp add: assms(1) bet_out out__acute) } thus ?thesis using \Per B A C \ Acute A B C\ assms(3) by blast next assume P2: "\ Col A B C" then have P3: "B \ C" using col_trivial_2 by auto obtain B' where P4: "Bet B A B' \ Cong A B' B A" using segment_construction by blast have P5: "\ Col B' A C" by (metis Col_def P2 P4 col_transitivity_2 cong_reverse_identity) then have P6: "B' \ A \ B' \ C" using not_col_distincts by blast then have P7: "A C B LtA C A B' \ A B C LtA C A B'" using P2 P4 l11_41 by auto then have P7A: "A C B LtA C A B'" by simp have P7B: "A B C LtA C A B'" by (simp add: P7) { assume "Per B A C" have "Acute A B C" by (metis Acute_def P4 P7B \Per B A C\ assms(1) bet_col col_per2__per col_trivial_3 l8_3 lta_right_comm) } { assume T1: "Obtuse B A C" then obtain a b c where T2: "Per a b c \ a b c LtA B A C" using Obtuse_def by blast then have T2A: "Per a b c" by simp have T2B: "a b c LtA B A C" by (simp add: T2) then have T3: "a b c LeA B A C \ \ a b c CongA B A C" by (simp add: LtA_def) then have T3A: "a b c LeA B A C" by simp have T3B: "\ a b c CongA B A C" by (simp add: T3) obtain P where T4: "P InAngle B A C \ a b c CongA B A P" using LeA_def T3 by blast then have T5: "Per B A P" using T4 T2 l11_17 by blast then have T6: "Per P A B" using l8_2 by blast have "Col A B B'" by (simp add: P4 bet_col col_permutation_4) then have "Per P A B'" using T6 assms(1) per_col by blast then have S3: "B A P CongA B' A P" using l8_2 P6 T5 T4 CongA_def assms(1) l11_16 by auto have "C A B' LtA P A B" proof - have S4: "B A P LeA B A C \ B' A C LeA B' A P" using P4 P6 assms(1) l11_36 by auto have S5: "C A B' LeA P A B" proof - have S6: "B A P LeA B A C" using T4 inangle__lea by auto have "B' A P CongA P A B" using S3 conga_left_comm not_conga_sym by blast thus ?thesis using P6 S4 S6 assms(2) conga_pseudo_refl l11_30 by auto qed { assume T10: "C A B' CongA P A B" have "Per B' A C" proof - have "B A P CongA B' A C" using T10 conga_comm conga_sym by blast thus ?thesis using T5 l11_17 by blast qed then have "Per C A B" using Col_cases P6 \Col A B B'\ l8_2 l8_3 by blast have "a b c CongA B A C" proof - have "a \ b" using T3A lea_distincts by auto have "c \ b" using T2B lta_distincts by blast have "Per B A C" using Per_cases \Per C A B\ by blast thus ?thesis using T2 \a \ b\ \c \ b\ assms(1) assms(2) l11_16 by auto qed then have "False" using T3B by blast } then have "\ C A B' CongA P A B" by blast thus ?thesis by (simp add: LtA_def S5) qed then have "A B C LtA B A P" by (meson P7 lta_right_comm lta_trans) then have "Acute A B C" using T5 using Acute_def by blast } thus ?thesis using \Per B A C \ Acute A B C\ assms(3) by blast qed lemma l11_43: assumes "A \ B" and "A \ C" and "Per B A C \ Obtuse B A C" shows "Acute A B C \ Acute A C B" using Per_perm assms(1) assms(2) assms(3) l11_43_aux obtuse_sym by blast lemma acute_lea_acute: assumes "Acute D E F" and "A B C LeA D E F" shows "Acute A B C" proof - obtain A' B' C' where P1: "Per A' B' C' \ D E F LtA A' B' C'" using Acute_def assms(1) by auto have P2: "A B C LeA A' B' C'" using LtA_def P1 assms(2) lea_trans by blast have "\ A B C CongA A' B' C'" by (meson LtA_def P1 assms(2) conga__lea456123 lea_asym lea_trans) then have "A B C LtA A' B' C'" by (simp add: LtA_def P2) thus ?thesis using Acute_def P1 by auto qed lemma lea_obtuse_obtuse: assumes "Obtuse D E F" and "D E F LeA A B C" shows "Obtuse A B C" proof - obtain A' B' C' where P1: "Per A' B' C' \ A' B' C' LtA D E F" using Obtuse_def assms(1) by auto then have P2: "A' B' C' LeA A B C" using LtA_def assms(2) lea_trans by blast have "\ A' B' C' CongA A B C" by (meson LtA_def P1 assms(2) conga__lea456123 lea_asym lea_trans) then have "A' B' C' LtA A B C" by (simp add: LtA_def P2) thus ?thesis using Obtuse_def P1 by auto qed lemma l11_44_1_a: assumes "A \ B" and "A \ C" and "Cong B A B C" shows "B A C CongA B C A" by (metis (no_types, opaque_lifting) Cong3_def assms(1) assms(2) assms(3) cong3_conga cong_inner_transitivity cong_pseudo_reflexivity) lemma l11_44_2_a: assumes "\ Col A B C" and "B A Lt B C" shows "B C A LtA B A C" proof - have T1: "A \ B" using assms(1) col_trivial_1 by auto have T3: "A \ C" using assms(1) col_trivial_3 by auto have "B A Le B C" by (simp add: assms(2) lt__le) then obtain C' where P1: "Bet B C' C \ Cong B A B C'" using assms(2) Le_def by blast have T5: "C \ C'" using P1 assms(2) cong__nlt by blast have T5A: "C' \ A" using Col_def Col_perm P1 assms(1) by blast then have T6: "C' InAngle B A C" using InAngle_def P1 T1 T3 out_trivial by auto have T7: "C' A C LtA A C' B \ C' C A LtA A C' B" proof - have W1: "\ Col C' C A" by (metis Col_def P1 T5 assms(1) col_transitivity_2) have W2: "Bet C C' B" using Bet_perm P1 by blast have "C' \ B" using P1 T1 cong_identity by blast thus ?thesis using l11_41 W1 W2 by simp qed have T90: "B A C' LtA B A C" proof - have T90A: "B A C' LeA B A C" by (simp add: T6 inangle__lea) have "B A C' CongA B A C'" using T1 T5A conga_refl by auto { assume "B A C' CongA B A C" then have R1: "A Out C' C" by (metis P1 T7 assms(1) bet_out conga_os__out lta_distincts not_col_permutation_4 out_one_side) have "B A OS C' C" by (metis Col_perm P1 T1 assms(1) bet_out cong_diff_2 out_one_side) then have "False" using Col_perm P1 T5 R1 bet_col col2__eq one_side_not_col123 out_col by blast } then have "\ B A C' CongA B A C" by blast thus ?thesis by (simp add: LtA_def T90A) qed have "B A C' CongA B C' A" using P1 T1 T5A l11_44_1_a by auto then have K2: "A C' B CongA B A C'" using conga_left_comm not_conga_sym by blast have "B C A LtA B A C'" proof - have K1: "B C A CongA B C A" using assms(1) conga_refl not_col_distincts by blast have "B C A LtA A C' B" proof - have "C' C A CongA B C A" proof - have K2: "C Out B C'" using P1 T5 bet_out_1 l6_6 by auto have "C Out A A" by (simp add: T3 out_trivial) thus ?thesis by (simp add: K2 out2__conga) qed have "A C' B CongA A C' B" using CongA_def K2 conga_refl by auto thus ?thesis using T7 \C' C A CongA B C A\ conga_preserves_lta by auto qed thus ?thesis using K1 K2 conga_preserves_lta by auto qed thus ?thesis using T90 lta_trans by blast qed lemma not_lta_and_conga: "\ ( A B C LtA D E F \ A B C CongA D E F)" by (simp add: LtA_def) lemma conga_sym_equiv: "A B C CongA A' B' C' \ A' B' C' CongA A B C" using not_conga_sym by blast lemma conga_dec: "A B C CongA D E F \ \ A B C CongA D E F" by auto lemma lta_not_conga: assumes "A B C LtA D E F" shows "\ A B C CongA D E F" using assms not_lta_and_conga by auto lemma lta__lea: assumes "A B C LtA D E F" shows "A B C LeA D E F" using LtA_def assms by auto lemma nlta: "\ A B C LtA A B C" using not_and_lta by blast lemma lea__nlta: assumes "A B C LeA D E F" shows "\ D E F LtA A B C" by (meson Tarski_neutral_dimensionless.lea_asym Tarski_neutral_dimensionless.not_lta_and_conga Tarski_neutral_dimensionless_axioms assms lta__lea) lemma lta__nlea: assumes "A B C LtA D E F" shows "\ D E F LeA A B C" using assms lea__nlta by blast lemma l11_44_1_b: assumes "\ Col A B C" and "B A C CongA B C A" shows "Cong B A B C" proof - have "B A Lt B C \ B A Gt B C \ Cong B A B C" by (simp add: or_lt_cong_gt) thus ?thesis by (meson Gt_def assms(1) assms(2) conga_sym l11_44_2_a not_col_permutation_3 not_lta_and_conga) qed lemma l11_44_2_b: assumes "B A C LtA B C A" shows "B C Lt B A" proof cases assume "Col A B C" thus ?thesis using Col_perm assms bet__lt1213 col_lta__bet lta_distincts by blast next assume P1: "\ Col A B C" then have P2: "A \ B" using col_trivial_1 by blast have P3: "A \ C" using P1 col_trivial_3 by auto have "B A Lt B C \ B A Gt B C \ Cong B A B C" by (simp add: or_lt_cong_gt) { assume "B A Lt B C" then have "B C Lt B A" using P1 assms l11_44_2_a not_and_lta by blast } { assume "B A Gt B C" then have "B C Lt B A" using Gt_def P1 assms l11_44_2_a not_and_lta by blast } { assume "Cong B A B C" then have "B A C CongA B C A" by (simp add: P2 P3 l11_44_1_a) then have "B C Lt B A" using assms not_lta_and_conga by blast } thus ?thesis by (meson P1 Tarski_neutral_dimensionless.not_and_lta Tarski_neutral_dimensionless_axioms \B A Gt B C \ B C Lt B A\ \B A Lt B C \ B A Gt B C \ Cong B A B C\ assms l11_44_2_a) qed lemma l11_44_1: assumes "\ Col A B C" shows "B A C CongA B C A \ Cong B A B C" using assms l11_44_1_a l11_44_1_b not_col_distincts by blast lemma l11_44_2: assumes "\ Col A B C" shows "B A C LtA B C A \ B C Lt B A" using assms l11_44_2_a l11_44_2_b not_col_permutation_3 by blast lemma l11_44_2bis: assumes "\ Col A B C" shows "B A C LeA B C A \ B C Le B A" proof - { assume P1: "B A C LeA B C A" { assume "B A Lt B C" then have "B C A LtA B A C" by (simp add: assms l11_44_2_a) then have "False" using P1 lta__nlea by auto } then have "\ B A Lt B C" by blast have "B C Le B A" using \\ B A Lt B C\ nle__lt by blast } { assume P2: "B C Le B A" have "B A C LeA B C A" proof cases assume "Cong B C B A" then have "B A C CongA B C A" by (metis assms conga_sym l11_44_1_a not_col_distincts) thus ?thesis by (simp add: conga__lea) next assume "\ Cong B C B A" then have "B A C LtA B C A" by (simp add: l11_44_2 assms Lt_def P2) thus ?thesis by (simp add: lta__lea) qed } thus ?thesis using \B A C LeA B C A \ B C Le B A\ by blast qed lemma l11_46: assumes "A \ B" and "B \ C" and "Per A B C \ Obtuse A B C" shows "B A Lt A C \ B C Lt A C" proof cases assume "Col A B C" thus ?thesis by (meson assms(1) assms(2) assms(3) bet__lt1213 bet__lt2313 col_obtuse__bet lt_left_comm per_not_col) next assume P1: "\ Col A B C" have P2: "A \ C" using P1 col_trivial_3 by auto have P3: "Acute B A C \ Acute B C A" using assms(1) assms(2) assms(3) l11_43 by auto then obtain A' B' C' where P4: "Per A' B' C' \ B C A LtA A' B' C'" using Acute_def P3 by auto { assume P5: "Per A B C" have P5A: "A C B CongA A C B" by (simp add: P2 assms(2) conga_refl) have S1: "A \ B" by (simp add: assms(1)) have S2: "B \ C" by (simp add: assms(2)) have S3: "A' \ B'" using P4 lta_distincts by blast have S4: "B' \ C'" using P4 lta_distincts by blast then have "A' B' C' CongA A B C" using l11_16 using S1 S2 S3 S4 P4 P5 by blast then have "A C B LtA A B C" using P5A P4 conga_preserves_lta lta_left_comm by blast } { assume "Obtuse A B C" obtain A'' B'' C'' where P6: "Per A'' B'' C'' \ A'' B'' C'' LtA A B C" using Obtuse_def \Obtuse A B C\ by auto have "B C A LtA A' B' C'" by (simp add: P4) then have P7: "A C B LtA A' B' C'" by (simp add: lta_left_comm) have "A' B' C' LtA A B C" proof - have U1: "A'' B'' C'' CongA A' B' C'" proof - have V2: "A'' \ B''" using P6 lta_distincts by blast have V3: "C'' \ B''" using P6 lta_distincts by blast have V5: "A' \ B'" using P7 lta_distincts by blast have "C' \ B'" using P4 lta_distincts by blast thus ?thesis using P6 V2 V3 P4 V5 by (simp add: l11_16) qed have U2: "A B C CongA A B C" using assms(1) assms(2) conga_refl by auto have U3: "A'' B'' C'' LtA A B C" by (simp add: P6) thus ?thesis using U1 U2 conga_preserves_lta by auto qed then have "A C B LtA A B C" using P7 lta_trans by blast } then have "A C B LtA A B C" using \Per A B C \ A C B LtA A B C\ assms(3) by blast then have "A B Lt A C" by (simp add: l11_44_2_b) then have "B A Lt A C" using Lt_cases by blast have "C A B LtA C B A" proof - obtain A' B' C' where U4: "Per A' B' C' \ B A C LtA A' B' C'" using Acute_def P3 by blast { assume "Per A B C" then have W3: "A' B' C' CongA C B A" using U4 assms(2) l11_16 l8_2 lta_distincts by blast have W2: "C A B CongA C A B" using P2 assms(1) conga_refl by auto have "C A B LtA A' B' C'" by (simp add: U4 lta_left_comm) then have "C A B LtA C B A" using W2 W3 conga_preserves_lta by blast } { assume "Obtuse A B C" then obtain A'' B'' C'' where W4: "Per A'' B'' C'' \ A'' B'' C'' LtA A B C" using Obtuse_def by auto have W5: "C A B LtA A' B' C'" by (simp add: U4 lta_left_comm) have "A' B' C' LtA C B A" proof - have W6: "A'' B'' C'' CongA A' B' C'" using l11_16 W4 U4 using lta_distincts by blast have "C B A CongA C B A" using assms(1) assms(2) conga_refl by auto thus ?thesis using W4 W6 conga_left_comm conga_preserves_lta by blast qed then have "C A B LtA C B A" using W5 lta_trans by blast } thus ?thesis using \Per A B C \ C A B LtA C B A\ assms(3) by blast qed then have "C B Lt C A" by (simp add: l11_44_2_b) then have "C B Lt A C" using Lt_cases by auto then have "B C Lt A C" using Lt_cases by blast thus ?thesis by (simp add: \B A Lt A C\) qed lemma l11_47: assumes "Per A C B" and "H PerpAt C H A B" shows "Bet A H B \ A \ H \ B \ H" proof - have P1: "Per C H A" using assms(2) perp_in_per_1 by auto have P2: "C H Perp A B" using assms(2) perp_in_perp by auto thus ?thesis proof cases assume "Col A C B" thus ?thesis by (metis P1 assms(1) assms(2) per_distinct_1 per_not_col perp_in_distinct perp_in_id) next assume P3: "\ Col A C B" have P4: "A \ H" by (metis P2 Per_perm Tarski_neutral_dimensionless.l8_7 Tarski_neutral_dimensionless_axioms assms(1) assms(2) col_trivial_1 perp_in_per_2 perp_not_col2) have P5: "Per C H B" using assms(2) perp_in_per_2 by auto have P6: "B \ H" using P1 P2 assms(1) l8_2 l8_7 perp_not_eq_1 by blast have P7: "H A Lt A C \ H C Lt A C" by (metis P1 P2 P4 l11_46 l8_2 perp_distinct) have P8: "C A Lt A B \ C B Lt A B" using P3 assms(1) l11_46 not_col_distincts by blast have P9: "H B Lt B C \ H C Lt B C" by (metis P2 P5 P6 Per_cases l11_46 perp_not_eq_1) have P10: "Bet A H B" proof - have T1: "Col A H B" using assms(2) col_permutation_5 perp_in_col by blast have T2: "A H Le A B" using P7 P8 by (meson lt_comm lt_transitivity nlt__le not_and_lt) have "H B Le A B" by (meson Lt_cases P8 P9 le_transitivity local.le_cases lt__nle) thus ?thesis using T1 T2 l5_12_b by blast qed thus ?thesis by (simp add: P4 P6) qed qed lemma l11_49: assumes "A B C CongA A' B' C'" and "Cong B A B' A'" and "Cong B C B' C'" shows "Cong A C A' C' \ (A \ C \ (B A C CongA B' A' C' \ B C A CongA B' C' A'))" proof - have T1:" Cong A C A' C'" using assms(1) assms(2) assms(3) cong2_conga_cong not_cong_2143 by blast { assume P1: "A \ C" have P2: "A \ B" using CongA_def assms(1) by blast have P3: "C \ B" using CongA_def assms(1) by blast have "B A C Cong3 B' A' C'" by (simp add: Cong3_def T1 assms(2) assms(3)) then have T2: "B A C CongA B' A' C'" using P1 P2 cong3_conga by auto have "B C A Cong3 B' C' A'" using Cong3_def T1 assms(2) assms(3) cong_3_swap_2 by blast then have "B C A CongA B' C' A'" using P1 P3 cong3_conga by auto then have "B A C CongA B' A' C' \ B C A CongA B' C' A'" using T2 by blast } thus ?thesis by (simp add: T1) qed lemma l11_50_1: assumes "\ Col A B C" and "B A C CongA B' A' C'" and "A B C CongA A' B' C'" and "Cong A B A' B'" shows "Cong A C A' C' \ Cong B C B' C' \ A C B CongA A' C' B'" proof - obtain C'' where P1: "B' Out C'' C' \ Cong B' C'' B C" by (metis Col_perm assms(1) assms(3) col_trivial_3 conga_diff56 l6_11_existence) have P2: "B' \ C''" using P1 out_diff1 by auto have P3: "\ Col A' B' C'" using assms(1) assms(3) ncol_conga_ncol by blast have P4: "\ Col A' B' C''" by (meson P1 P2 P3 col_transitivity_1 not_col_permutation_2 out_col) have P5: "Cong A C A' C''" proof - have Q1: "B Out A A" using assms(1) not_col_distincts out_trivial by auto have Q2: "B Out C C" using assms(1) col_trivial_2 out_trivial by force have Q3: "B' Out A' A'" using P3 not_col_distincts out_trivial by auto have Q5: "Cong B A B' A'" using assms(4) not_cong_2143 by blast have "Cong B C B' C''" using P1 not_cong_3412 by blast thus ?thesis using l11_4_1 P1 Q1 Q2 Q3 Q5 assms(3) by blast qed have P6: "B A C Cong3 B' A' C''" using Cong3_def Cong_perm P1 P5 assms(4) by blast have P7: "B A C CongA B' A' C''" by (metis P6 assms(1) cong3_conga not_col_distincts) have P8: "B' A' C' CongA B' A' C''" by (meson P7 assms(2) conga_sym conga_trans) have "B' A' OS C' C''" using Col_perm Out_cases P1 P3 out_one_side by blast then have "A' Out C' C''" using P8 conga_os__out by auto then have "Col A' C' C''" using out_col by auto then have P9: "C' = C''" using Col_perm P1 out_col P3 col_transitivity_1 by blast have T1: "Cong A C A' C'" by (simp add: P5 P9) have T2: "Cong B C B' C'" using Cong_perm P1 P9 by blast then have "A C B CongA A' C' B'" using T1 assms(1) assms(2) assms(4) col_trivial_2 l11_49 by blast thus ?thesis using T1 T2 by blast qed lemma l11_50_2: assumes "\ Col A B C" and "B C A CongA B' C' A'" and "A B C CongA A' B' C'" and "Cong A B A' B'" shows "Cong A C A' C' \ Cong B C B' C' \ C A B CongA C' A' B'" proof - have P1: "A \ B" using assms(1) col_trivial_1 by auto have P2: "B \ C" using assms(1) col_trivial_2 by auto have P3: "A' \ B'" using P1 assms(4) cong_diff by blast have P4: "B' \ C'" using assms(2) conga_diff45 by auto then obtain C'' where P5: "B' Out C'' C' \ Cong B' C'' B C" using P2 l6_11_existence by presburger have P5BIS: "B' \ C''" using P5 out_diff1 by auto have P5A: "Col B' C'' C'" using P5 out_col by auto have P6: "\ Col A' B' C'" using assms(1) assms(3) ncol_conga_ncol by blast { assume "Col A' B' C''" then have "Col B' C'' A'" using not_col_permutation_2 by blast then have "Col B' C' A'" using col_transitivity_1 P5BIS P5A by blast then have "Col A' B' C'" using Col_perm by blast then have False using P6 by auto } then have P7: "\ Col A' B' C''" by blast have P8: "Cong A C A' C''" proof - have "B Out A A" by (simp add: P1 out_trivial) have K1: "B Out C C" using P2 out_trivial by auto have K2: "B' Out A' A'" using P3 out_trivial by auto have "Cong B A B' A'" by (simp add: Cong_perm assms(4)) have "Cong B C B' C''" using Cong_perm P5 by blast thus ?thesis using P5 \Cong B A B' A'\ P1 out_trivial K1 K2 assms(3) l11_4_1 by blast qed have P9: "B C A Cong3 B' C'' A'" using Cong3_def Cong_perm P5 P8 assms(4) by blast then have P10: "B C A CongA B' C'' A'" using assms(1) cong3_conga not_col_distincts by auto have P11: "B' C' A' CongA B' C'' A'" using P9 assms(2) cong3_conga2 conga_sym by blast show ?thesis proof cases assume L1: "C' = C''" then have L2: "Cong A C A' C'" by (simp add: P8) have L3: "Cong B C B' C'" using Cong_perm L1 P5 by blast have "C A B Cong3 C' A' B'" by (simp add: L1 P9 cong_3_swap cong_3_swap_2) then have "C A B CongA C' A' B'" by (metis CongA_def P1 assms(2) cong3_conga) thus ?thesis using L2 L3 by auto next assume R1: "C' \ C''" have R1A: "\ Col C'' C' A'" by (metis P5A P7 R1 col_permutation_2 col_trivial_2 colx) have R1B: "Bet B' C'' C' \ Bet B' C' C''" using Out_def P5 by auto { assume S1: "Bet B' C'' C'" then have S2: "C'' A' C' LtA A' C'' B' \ C'' C' A' LtA A' C'' B'" using P5BIS R1A between_symmetry l11_41 by blast have "B' C' A' CongA C'' C' A'" by (metis P11 R1 Tarski_neutral_dimensionless.conga_comm Tarski_neutral_dimensionless_axioms S1 bet_out_1 conga_diff45 not_conga_sym out2__conga out_trivial) then have "B' C' A' LtA A' C'' B'" by (meson P11 Tarski_neutral_dimensionless.conga_right_comm Tarski_neutral_dimensionless.not_conga Tarski_neutral_dimensionless.not_conga_sym Tarski_neutral_dimensionless_axioms S2 not_lta_and_conga) then have "Cong A C A' C' \ Cong B C B' C'" by (meson P11 Tarski_neutral_dimensionless.conga_right_comm Tarski_neutral_dimensionless_axioms not_lta_and_conga) } { assume Z1: "Bet B' C' C''" have Z2: "\ Col C' C'' A'" by (simp add: R1A not_col_permutation_4) have Z3: "C'' Out C' B'" by (simp add: R1 Z1 bet_out_1) have Z4: "C'' Out A' A'" using P7 not_col_distincts out_trivial by blast then have Z4A: "B' C'' A' CongA C' C'' A'" by (simp add: Z3 out2__conga) have Z4B: "B' C'' A' LtA A' C' B'" proof - have Z5: "C' C'' A' CongA B' C'' A'" using Z4A not_conga_sym by blast have Z6: "A' C' B' CongA A' C' B'" using P11 P4 conga_diff2 conga_refl by blast have "C' C'' A' LtA A' C' B'" using P4 Z1 Z2 between_symmetry l11_41 by blast thus ?thesis using Z5 Z6 conga_preserves_lta by auto qed have "B' C'' A' CongA B' C' A'" using P11 not_conga_sym by blast then have "Cong A C A' C' \ Cong B C B' C'" by (meson Tarski_neutral_dimensionless.conga_right_comm Tarski_neutral_dimensionless_axioms Z4B not_lta_and_conga) } then have R2: "Cong A C A' C' \ Cong B C B' C'" using R1B \Bet B' C'' C' \ Cong A C A' C' \ Cong B C B' C'\ by blast then have "C A B CongA C' A' B'" using P1 assms(2) l11_49 not_cong_2143 by blast thus ?thesis using R2 by auto qed qed lemma l11_51: assumes "A \ B" and "A \ C" and "B \ C" and "Cong A B A' B'" and "Cong A C A' C'" and "Cong B C B' C'" shows "B A C CongA B' A' C' \ A B C CongA A' B' C' \ B C A CongA B' C' A'" proof - have "B A C Cong3 B' A' C' \ A B C Cong3 A' B' C' \ B C A Cong3 B' C' A'" using Cong3_def Cong_perm assms(4) assms(5) assms(6) by blast thus ?thesis using assms(1) assms(2) assms(3) cong3_conga by auto qed lemma conga_distinct: assumes "A B C CongA D E F" shows "A \ B \ C \ B \ D \ E \ F \ E" using CongA_def assms by auto lemma l11_52: assumes "A B C CongA A' B' C'" and "Cong A C A' C'" and "Cong B C B' C'" and "B C Le A C" shows "Cong B A B' A' \ B A C CongA B' A' C' \ B C A CongA B' C' A'" proof - have P1: "A \ B" using CongA_def assms(1) by blast have P2: "C \ B" using CongA_def assms(1) by blast have P3: "A' \ B'" using CongA_def assms(1) by blast have P4: "C' \ B'" using assms(1) conga_diff56 by auto have P5: "Cong B A B' A'" proof cases assume P6: "Col A B C" then have P7: "Bet A B C \ Bet B C A \ Bet C A B" using Col_def by blast { assume P8: "Bet A B C" then have "Bet A' B' C'" using assms(1) bet_conga__bet by blast then have "Cong B A B' A'" using P8 assms(2) assms(3) l4_3 not_cong_2143 by blast } { assume P9: "Bet B C A" then have P10: "B' Out A' C'" using Out_cases P2 assms(1) bet_out l11_21_a by blast then have P11: "Bet B' A' C' \ Bet B' C' A'" by (simp add: Out_def) { assume "Bet B' A' C'" then have "Cong B A B' A'" using P3 assms(2) assms(3) assms(4) bet_le_eq l5_6 by blast } { assume "Bet B' C' A'" then have "Cong B A B' A'" using Cong_perm P9 assms(2) assms(3) l2_11_b by blast } then have "Cong B A B' A'" using P11 \Bet B' A' C' \ Cong B A B' A'\ by blast } { assume "Bet C A B" then have "Cong B A B' A'" using P1 assms(4) bet_le_eq between_symmetry by blast } thus ?thesis using P7 \Bet A B C \ Cong B A B' A'\ \Bet B C A \ Cong B A B' A'\ by blast next assume Z1: "\ Col A B C" obtain A'' where Z2: "B' Out A'' A' \ Cong B' A'' B A" using P1 P3 l6_11_existence by force then have Z3: "A' B' C' CongA A'' B' C'" by (simp add: P4 out2__conga out_trivial) have Z4: "A B C CongA A'' B' C'" using Z3 assms(1) not_conga by blast have Z5: "Cong A'' C' A C" using Z2 Z4 assms(3) cong2_conga_cong cong_4321 cong_symmetry by blast have Z6: "A'' B' C' Cong3 A B C" using Cong3_def Cong_perm Z2 Z5 assms(3) by blast have Z7: "Cong A'' C' A' C'" using Z5 assms(2) cong_transitivity by blast have Z8: "\ Col A' B' C'" by (metis Z1 assms(1) ncol_conga_ncol) then have Z9: "\ Col A'' B' C'" by (metis Z2 col_transitivity_1 not_col_permutation_4 out_col out_diff1) { assume Z9A: "A'' \ A'" have Z10: "Bet B' A'' A' \ Bet B' A' A''" using Out_def Z2 by auto { assume Z11: "Bet B' A'' A'" have Z12: "A'' C' B' LtA C' A'' A' \ A'' B' C' LtA C' A'' A'" by (simp add: Z11 Z9 Z9A l11_41) have Z13: "Cong A' C' A'' C'" using Cong_perm Z7 by blast have Z14: "\ Col A'' C' A'" by (metis Col_def Z11 Z9 Z9A col_transitivity_1) have Z15: "C' A'' A' CongA C' A' A'' \ Cong C' A'' C' A'" by (simp add: Z14 l11_44_1) have Z16: "Cong C' A' C' A''" using Cong_perm Z7 by blast then have Z17: "Cong C' A'' C' A'" using Cong_perm by blast then have Z18: "C' A'' A' CongA C' A' A''" by (simp add: Z15) have Z19: "\ Col B' C' A''" using Col_perm Z9 by blast have Z20: "B' A' C' CongA A'' A' C'" by (metis Tarski_neutral_dimensionless.col_conga_col Tarski_neutral_dimensionless_axioms Z11 Z3 Z9 Z9A bet_out_1 col_trivial_3 out2__conga out_trivial) have Z21: "\ Col B' C' A'" using Col_perm Z8 by blast then have Z22: "C' B' A' LtA C' A' B' \ C' A' Lt C' B'" by (simp add: l11_44_2) have "A'' B' C' CongA C' B' A'" using Z3 conga_right_comm not_conga_sym by blast then have U1: "C' B' A' LtA C' A' B'" proof - have f1: "\p pa pb pc pd pe pf pg ph pi pj pk pl pm. \ Tarski_neutral_dimensionless p pa \ \ Tarski_neutral_dimensionless.CongA p pa (pb::'p) pc pd pe pf pg \ \ Tarski_neutral_dimensionless.CongA p pa ph pi pj pk pl pm \ \ Tarski_neutral_dimensionless.LtA p pa pb pc pd ph pi pj \ Tarski_neutral_dimensionless.LtA p pa pe pf pg pk pl pm" by (simp add: Tarski_neutral_dimensionless.conga_preserves_lta) have f2: "C' A'' A' CongA C' A' A''" by (metis Z15 Z17) have f3: "\p pa pb pc pd pe pf pg. \ Tarski_neutral_dimensionless p pa \ \ Tarski_neutral_dimensionless.CongA p pa (pb::'p) pc pd pe pf pg \ Tarski_neutral_dimensionless.CongA p pa pe pf pg pb pc pd" by (metis (no_types) Tarski_neutral_dimensionless.conga_sym) then have "\ C' B' A' LtA C' A'' A' \ A'' B' C' LtA C' A' A''" using f2 f1 by (meson Tarski_neutral_dimensionless_axioms \A'' B' C' CongA C' B' A'\) then have "C' B' A' LtA C' A' B' \ A'' B' C' LtA A'' A' C' \ A'' = B'" using f2 f1 by (metis (no_types) Tarski_neutral_dimensionless.conga_refl Tarski_neutral_dimensionless_axioms Z12 \A'' B' C' CongA C' B' A'\ lta_right_comm) thus ?thesis using f3 f2 f1 by (metis (no_types) Tarski_neutral_dimensionless_axioms Z12 Z20 \A'' B' C' CongA C' B' A'\ lta_right_comm) qed then have Z23: "C' A' Lt C' B'" using Z22 by auto have Z24: "C' A'' Lt C' B'" using Z16 Z23 cong2_lt__lt cong_reflexivity by blast have Z25: "C A Le C B" proof - have Z26: "Cong C' A'' C A" using Z5 not_cong_2143 by blast have "Cong C' B' C B" using assms(3) not_cong_4321 by blast thus ?thesis using l5_6 Z24 Z26 lt__le by blast qed then have Z27: "Cong C A C B" by (simp add: assms(4) le_anti_symmetry le_comm) have "Cong C' A'' C' B'" by (metis Cong_perm Z13 Z27 assms(2) assms(3) cong_transitivity) then have "False" using Z24 cong__nlt by blast then have "Cong B A B' A'" by simp } { assume W1: "Bet B' A' A''" have W2: "A' \ A''" using Z9A by auto have W3: "A' C' B' LtA C' A' A'' \ A' B' C' LtA C' A' A''" using W1 Z8 Z9A l11_41 by blast have W4: "Cong A' C' A'' C'" using Z7 not_cong_3412 by blast have "\ Col A'' C' A'" by (metis Col_def W1 Z8 Z9A col_transitivity_1) then have W6: "C' A'' A' CongA C' A' A'' \ Cong C' A'' C' A'" using l11_44_1 by auto have W7: "Cong C' A' C' A''" using Z7 not_cong_4321 by blast then have W8: "Cong C' A'' C' A'" using W4 not_cong_4321 by blast have W9: "\ Col B' C' A''" by (simp add: Z9 not_col_permutation_1) have W10: "B' A'' C' CongA A' A'' C'" by (metis Tarski_neutral_dimensionless.Out_def Tarski_neutral_dimensionless_axioms W1 Z9 Z9A bet_out_1 between_trivial not_col_distincts out2__conga) have W12: "C' B' A'' LtA C' A'' B' \ C' A'' Lt C' B'" by (simp add: W9 l11_44_2) have W12A: "C' B' A'' LtA C' A'' B'" proof - have V1: "A' B' C' CongA C' B' A''" by (simp add: Z3 conga_right_comm) have "A' A'' C' CongA B' A'' C'" by (metis Tarski_neutral_dimensionless.Out_def Tarski_neutral_dimensionless_axioms W1 \\ Col A'' C' A'\ between_equality_2 not_col_distincts or_bet_out out2__conga out_col) then have "C' A' A'' CongA C' A'' B'" by (meson W6 W8 conga_left_comm not_conga not_conga_sym) thus ?thesis using W3 V1 conga_preserves_lta by auto qed then have "C' A'' Lt C' B'" using W12 by auto then have W14: "C' A' Lt C' B'" using W8 cong2_lt__lt cong_reflexivity by blast have W15: "C A Le C B" proof - have Q1: "C' A'' Le C' B'" using W12 W12A lt__le by blast have Q2: "Cong C' A'' C A" using Z5 not_cong_2143 by blast have "Cong C' B' C B" using assms(3) not_cong_4321 by blast thus ?thesis using Q1 Q2 l5_6 by blast qed have "C B Le C A" by (simp add: assms(4) le_comm) then have "Cong C A C B" by (simp add: W15 le_anti_symmetry) then have "Cong C' A' C' B'" by (metis Cong_perm assms(2) assms(3) cong_inner_transitivity) then have "False" using W14 cong__nlt by blast then have "Cong B A B' A'" by simp } then have "Cong B A B' A'" using Z10 \Bet B' A'' A' \ Cong B A B' A'\ by blast } { assume "A'' = A'" then have "Cong B A B' A'" using Z2 not_cong_3412 by blast } thus ?thesis using \A'' \ A' \ Cong B A B' A'\ by blast qed have P6: "A B C Cong3 A' B' C'" using Cong3_def Cong_perm P5 assms(2) assms(3) by blast thus ?thesis using P2 P5 assms(1) assms(3) assms(4) l11_49 le_zero by blast qed lemma l11_53: assumes "Per D C B" and "C \ D" and "A \ B" and "B \ C" and "Bet A B C" shows "C A D LtA C B D \ B D Lt A D" proof - have P1: "C \ A" using assms(3) assms(5) between_identity by blast have P2: "\ Col B A D" by (smt assms(1) assms(2) assms(3) assms(4) assms(5) bet_col bet_col1 col3 col_permutation_4 l8_9) have P3: "A \ D" using P2 col_trivial_2 by blast have P4: "C A D LtA C B D" proof - have P4A: "B D A LtA D B C \ B A D LtA D B C" by (simp add: P2 assms(4) assms(5) l11_41) have P4AA:"A Out B C" using assms(3) assms(5) bet_out by auto have "A Out D D" using P3 out_trivial by auto then have P4B: "C A D CongA B A D" using P4AA by (simp add: out2__conga) then have P4C: "B A D CongA C A D" by (simp add: P4B conga_sym) have "D B C CongA C B D" using assms(1) assms(4) conga_pseudo_refl per_distinct_1 by auto thus ?thesis using P4A P4C conga_preserves_lta by blast qed obtain B' where P5: "C Midpoint B B' \ Cong D B D B'" using Per_def assms(1) by auto have K2: "A \ B'" using Bet_cases P5 assms(4) assms(5) between_equality_2 midpoint_bet by blast { assume "Col B D B'" then have "Col B A D" by (metis Col_cases P5 assms(1) assms(2) assms(4) col2__eq midpoint_col midpoint_distinct_2 per_not_col) then have "False" by (simp add: P2) } then have P6: "\ Col B D B'" by blast then have "D B B' CongA D B' B \ Cong D B D B'" by (simp add: l11_44_1) then have "D B B' CongA D B' B" using P5 by simp { assume K1: "Col A D B'" have "Col B' A B" using Col_def P5 assms(4) assms(5) midpoint_bet outer_transitivity_between by blast then have "Col B' B D" using K1 K2 Col_perm col_transitivity_2 by blast then have "Col B D B'" using Col_perm by blast then have "False" by (simp add: P6) } then have K3B: "\ Col A D B'" by blast then have K4: "D A B' LtA D B' A \ D B' Lt D A" by (simp add: l11_44_2) have K4A: "C A D LtA C B' D" by (metis Midpoint_def P1 P3 P4 P5 P5 P6 assms(2) assms(4) col_trivial_1 cong_reflexivity conga_preserves_lta conga_refl l11_51 not_cong_2134) have "D B' Lt D A" proof - have "D A B' LtA D B' A" proof - have K5A: "A Out D D" using P3 out_trivial by auto have K5AA: "A Out B' C" by (smt K2 Out_def P1 P5 assms(4) assms(5) midpoint_bet outer_transitivity_between2) then have K5: "D A C CongA D A B'" by (simp add: K5A out2__conga) have K6A: "B' Out D D" using K3B not_col_distincts out_trivial by blast have "B' Out A C" by (smt P5 K5AA assms(4) assms(5) between_equality_2 l6_4_2 midpoint_bet midpoint_distinct_2 out_col outer_transitivity_between2) then have K6: "D B' C CongA D B' A" by (simp add: K6A out2__conga) have "D A C LtA D B' C" by (simp add: K4A lta_comm) thus ?thesis using K5 K6 conga_preserves_lta by auto qed thus ?thesis by (simp add: K4) qed thus ?thesis using P4 P5 cong2_lt__lt cong_pseudo_reflexivity not_cong_4312 by blast qed lemma cong2_conga_obtuse__cong_conga2: assumes "Obtuse A B C" and "A B C CongA A' B' C'" and "Cong A C A' C'" and "Cong B C B' C'" shows "Cong B A B' A' \ B A C CongA B' A' C' \ B C A CongA B' C' A'" proof - have "B C Le A C" proof cases assume "Col A B C" thus ?thesis by (simp add: assms(1) col_obtuse__bet l5_12_a) next assume "\ Col A B C" thus ?thesis using l11_46 assms(1) lt__le not_col_distincts by auto qed thus ?thesis using l11_52 assms(2) assms(3) assms(4) by blast qed lemma cong2_per2__cong_conga2: assumes "A \ B" and "B \ C" and "Per A B C" and "Per A' B' C'" and "Cong A C A' C'" and "Cong B C B' C'" shows "Cong B A B' A' \ B A C CongA B' A' C' \ B C A CongA B' C' A'" proof - have P1: "B C Le A C \ \ Cong B C A C" using assms(1) assms(2) assms(3) cong__nlt l11_46 lt__le by blast then have "A B C CongA A' B' C'" using assms(2) assms(3) assms(4) assms(5) assms(6) cong_diff cong_inner_transitivity cong_symmetry l11_16 by blast thus ?thesis using P1 assms(5) assms(6) l11_52 by blast qed lemma cong2_per2__cong: assumes "Per A B C" and "Per A' B' C'" and "Cong A C A' C'" and "Cong B C B' C'" shows "Cong B A B' A'" proof cases assume "B = C" thus ?thesis using assms(3) assms(4) cong_reverse_identity not_cong_2143 by blast next assume "B \ C" show ?thesis proof cases assume "A = B" thus ?thesis proof - have "Cong A C B' C'" using \A = B\ assms(4) by blast then have "B' = A'" by (meson Cong3_def Per_perm assms(2) assms(3) cong_inner_transitivity cong_pseudo_reflexivity l8_10 l8_7) thus ?thesis using \A = B\ cong_trivial_identity by blast qed next assume "A \ B" show ?thesis proof cases assume "A' = B'" thus ?thesis by (metis Cong3_def Per_perm \A \ B\ assms(1) assms(3) assms(4) cong_inner_transitivity cong_pseudo_reflexivity l8_10 l8_7) next assume "A' \ B'" thus ?thesis using cong2_per2__cong_conga2 \A \ B\ \B \ C\ assms(1) assms(2) assms(3) assms(4) by blast qed qed qed lemma cong2_per2__cong_3: assumes "Per A B C" "Per A' B' C'" and "Cong A C A' C'" and "Cong B C B' C'" shows "A B C Cong3 A' B' C'" by (metis Tarski_neutral_dimensionless.Cong3_def Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) assms(4) cong2_per2__cong cong_3_swap) lemma cong_lt_per2__lt: assumes "Per A B C" and "Per A' B' C'" and "Cong A B A' B'" and "B C Lt B' C'" shows "A C Lt A' C'" proof cases assume "A = B" thus ?thesis using assms(3) assms(4) cong_reverse_identity by blast next assume "A \ B" show ?thesis proof cases assume "B = C" thus ?thesis by (smt assms(2) assms(3) assms(4) cong2_lt__lt cong_4312 cong_diff cong_reflexivity l11_46 lt_diff) next assume P0: "B \ C" have "B C Lt B' C'" by (simp add: assms(4)) then have R1: "B C Le B' C' \ \ Cong B C B' C'" by (simp add: Lt_def) then obtain C0 where P1: "Bet B' C0 C' \ Cong B C B' C0" using Le_def by auto then have P2: "Per A' B' C0" by (metis Col_def Per_cases assms(2) bet_out_1 col_col_per_per col_trivial_1 l8_5 out_diff2) have "C0 A' Lt C' A'" using l11_53 by (metis P1 P2 R1 P0 bet__lt2313 between_symmetry cong_diff) then have P3: "A' C0 Lt A' C'" using Lt_cases by blast have P4: "Cong A' C0 A C" using P1 P2 assms(1) assms(3) l10_12 not_cong_3412 by blast have "Cong A' C' A' C'" by (simp add: cong_reflexivity) thus ?thesis using cong2_lt__lt P3 P4 by blast qed qed lemma cong_le_per2__le: assumes "Per A B C" and "Per A' B' C'" and "Cong A B A' B'" and "B C Le B' C'" shows "A C Le A' C'" proof cases assume "Cong B C B' C'" thus ?thesis using assms(1) assms(2) assms(3) cong__le l10_12 by blast next assume "\ Cong B C B' C'" then have "B C Lt B' C'" using Lt_def assms(4) by blast thus ?thesis using assms(1) assms(2) assms(3) cong_lt_per2__lt lt__le by auto qed lemma lt2_per2__lt: assumes "Per A B C" and "Per A' B' C'" and "A B Lt A' B'" and "B C Lt B' C'" shows "A C Lt A' C'" proof - have P2: "B A Lt B' A'" by (simp add: assms(3) lt_comm) have P3: "B C Le B' C' \ \ Cong B C B' C'" using assms(4) cong__nlt lt__le by auto then obtain C0 where P4: "Bet B' C0 C' \ Cong B C B' C0" using Le_def by auto have P4A: "B' \ C'" using assms(4) lt_diff by auto have "Col B' C' C0" using P4 bet_col not_col_permutation_5 by blast then have P5: "Per A' B' C0" using assms(2) P4A per_col by blast have P6: "A C Lt A' C0" by (meson P2 P4 P5 assms(1) cong_lt_per2__lt l8_2 lt_comm not_cong_2143) have "B' C0 Lt B' C'" by (metis P4 assms(4) bet__lt1213 cong__nlt) then have "A' C0 Lt A' C'" using P5 assms(2) cong_lt_per2__lt cong_reflexivity by blast thus ?thesis using P6 lt_transitivity by blast qed lemma le_lt_per2__lt: assumes "Per A B C" and "Per A' B' C'" and "A B Le A' B'" and "B C Lt B' C'" shows "A C Lt A' C'" using Lt_def assms(1) assms(2) assms(3) assms(4) cong_lt_per2__lt lt2_per2__lt by blast lemma le2_per2__le: assumes "Per A B C" and "Per A' B' C'" and "A B Le A' B'" and "B C Le B' C'" shows "A C Le A' C'" proof cases assume "Cong B C B' C'" thus ?thesis by (meson Per_cases Tarski_neutral_dimensionless.cong_le_per2__le Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) le_comm not_cong_2143) next assume "\ Cong B C B' C'" then have "B C Lt B' C'" by (simp add: Lt_def assms(4)) thus ?thesis using assms(1) assms(2) assms(3) le_lt_per2__lt lt__le by blast qed lemma cong_lt_per2__lt_1: assumes "Per A B C" and "Per A' B' C'" and "A B Lt A' B'" and "Cong A C A' C'" shows "B' C' Lt B C" by (meson Gt_def assms(1) assms(2) assms(3) assms(4) cong2_per2__cong cong_4321 cong__nlt cong_symmetry lt2_per2__lt or_lt_cong_gt) lemma symmetry_preserves_conga: assumes "A \ B" and "C \ B" and "M Midpoint A A'" and "M Midpoint B B'" and "M Midpoint C C'" shows "A B C CongA A' B' C'" by (metis Mid_perm assms(1) assms(2) assms(3) assms(4) assms(5) conga_trivial_1 l11_51 l7_13 symmetric_point_uniqueness) lemma l11_57: assumes "A A' OS B B'" and "Per B A A'" and "Per B' A' A" and "A A' OS C C'" and "Per C A A'" and "Per C' A' A" shows "B A C CongA B' A' C'" proof - obtain M where P1: "M Midpoint A A'" using midpoint_existence by auto obtain B'' where P2: "M Midpoint B B''" using symmetric_point_construction by auto obtain C'' where P3: "M Midpoint C C''" using symmetric_point_construction by auto have P4: "\ Col A A' B" using assms(1) col123__nos by auto have P5: "\ Col A A' C" using assms(4) col123__nos by auto have P6: "B A C CongA B'' A' C''" by (metis P1 P2 P3 assms(1) assms(4) os_distincts symmetry_preserves_conga) have "B'' A' C'' CongA B' A' C'" proof - have "B \ M" using P1 P4 midpoint_col not_col_permutation_2 by blast then have P7: "\ Col B'' A A'" using Mid_cases P1 P2 P4 mid_preserves_col not_col_permutation_3 by blast have K3: "Bet B'' A' B'" proof - have "Per B'' A' A" using P1 P2 assms(2) per_mid_per by blast have "Col B B'' M \ Col A A' M" using P1 P2 midpoint_col not_col_permutation_2 by blast then have "Coplanar B A A' B''" using Coplanar_def by auto then have "Coplanar A B' B'' A'" by (meson assms(1) between_trivial2 coplanar_trans_1 ncoplanar_perm_4 ncoplanar_perm_8 one_side_chara os__coplanar) then have P8: "Col B' B'' A'" using cop_per2__col P1 P2 P7 assms(2) assms(3) not_col_distincts per_mid_per by blast have "A A' TS B B''" using P1 P2 P4 mid_two_sides by auto then have "A' A TS B'' B'" using assms(1) invert_two_sides l9_2 l9_8_2 by blast thus ?thesis using Col_cases P8 col_two_sides_bet by blast qed have "\ Col C'' A A'" by (smt Col_def P1 P3 P5 l7_15 l7_2 not_col_permutation_5) have "Bet C'' A' C'" proof - have Z2: "Col C' C'' A'" proof - have "Col C C'' M \ Col A A' M" using P1 P3 col_permutation_1 midpoint_col by blast then have "Coplanar C A A' C''" using Coplanar_def by blast then have Z1: "Coplanar A C' C'' A'" by (meson assms(4) between_trivial2 coplanar_trans_1 ncoplanar_perm_4 ncoplanar_perm_8 one_side_chara os__coplanar) have "Per C'' A' A" using P1 P3 assms(5) per_mid_per by blast thus ?thesis using Z1 P5 assms(6) col_trivial_1 cop_per2__col by blast qed have "A A' TS C C''" using P1 P3 P5 mid_two_sides by auto then have "A' A TS C'' C'" using assms(4) invert_two_sides l9_2 l9_8_2 by blast thus ?thesis using Col_cases Z2 col_two_sides_bet by blast qed thus ?thesis by (metis P6 K3 assms(1) assms(4) conga_diff45 conga_diff56 l11_14 os_distincts) qed thus ?thesis using P6 conga_trans by blast qed lemma cop3_orth_at__orth_at: assumes "\ Col D E F" and "Coplanar A B C D" and "Coplanar A B C E" and "Coplanar A B C F" and "X OrthAt A B C U V" shows "X OrthAt D E F U V" proof - have P1: "\ Col A B C \ Coplanar A B C X" using OrthAt_def assms(5) by blast then have P2: "Coplanar D E F X" using assms(2) assms(3) assms(4) coplanar_pseudo_trans by blast { fix M assume "Coplanar A B C M" then have "Coplanar D E F M" using P1 assms(2) assms(3) assms(4) coplanar_pseudo_trans by blast } have T1: "U \ V" using OrthAt_def assms(5) by blast have T2: "Col U V X" using OrthAt_def assms(5) by auto { fix P Q assume P7: "Coplanar D E F P \ Col U V Q" then have "Coplanar A B C P" by (meson \\M. Coplanar A B C M \ Coplanar D E F M\ assms(1) assms(2) assms(3) assms(4) l9_30) then have "Per P X Q" using P7 OrthAt_def assms(5) by blast } thus ?thesis using assms(1) by (simp add: OrthAt_def P2 T1 T2) qed lemma col2_orth_at__orth_at: assumes "U \ V" and "Col P Q U" and "Col P Q V" and "X OrthAt A B C P Q" shows "X OrthAt A B C U V" proof - have "Col P Q X" using OrthAt_def assms(4) by auto then have "Col U V X" by (metis OrthAt_def assms(2) assms(3) assms(4) col3) thus ?thesis using OrthAt_def assms(1) assms(2) assms(3) assms(4) colx by presburger qed lemma col_orth_at__orth_at: assumes "U \ W" and "Col U V W" and "X OrthAt A B C U V" shows "X OrthAt A B C U W" using assms(1) assms(2) assms(3) col2_orth_at__orth_at col_trivial_3 by blast lemma orth_at_symmetry: assumes "X OrthAt A B C U V" shows "X OrthAt A B C V U" by (metis assms col2_orth_at__orth_at col_trivial_2 col_trivial_3) lemma orth_at_distincts: assumes "X OrthAt A B C U V" shows "A \ B \ B \ C \ A \ C \ U \ V" using OrthAt_def assms not_col_distincts by fastforce lemma orth_at_chara: "X OrthAt A B C X P \ (\ Col A B C \ X \ P \ Coplanar A B C X \ (\ D.(Coplanar A B C D \ Per D X P)))" proof - { assume "X OrthAt A B C X P" then have "\ Col A B C \ X \ P \ Coplanar A B C X \ (\ D.(Coplanar A B C D \ Per D X P))" using OrthAt_def col_trivial_2 by auto } { assume T1: "\ Col A B C \ X \ P \ Coplanar A B C X \ (\ D.(Coplanar A B C D \ Per D X P))" { fix P0 Q assume "Coplanar A B C P0 \ Col X P Q" then have "Per P0 X Q" using T1 OrthAt_def per_col by auto } then have "X OrthAt A B C X P" by (simp add: T1 \\Q P0. Coplanar A B C P0 \ Col X P Q \ Per P0 X Q\ Tarski_neutral_dimensionless.OrthAt_def Tarski_neutral_dimensionless_axioms col_trivial_3) } thus ?thesis using \X OrthAt A B C X P \ \ Col A B C \ X \ P \ Coplanar A B C X \ (\D. Coplanar A B C D \ Per D X P)\ by blast qed lemma cop3_orth__orth: 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 Orth U V" shows "D E F Orth U V" using Orth_def assms(1) assms(2) assms(3) assms(4) assms(5) cop3_orth_at__orth_at by blast lemma col2_orth__orth: assumes "U \ V" and "Col P Q U" and "Col P Q V" and "A B C Orth P Q" shows "A B C Orth U V" by (meson Orth_def Tarski_neutral_dimensionless.col2_orth_at__orth_at Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) assms(4)) lemma col_orth__orth: assumes "U \ W" and "Col U V W" and "A B C Orth U V" shows "A B C Orth U W" by (meson assms(1) assms(2) assms(3) col2_orth__orth col_trivial_3) lemma orth_symmetry: assumes "A B C Orth U V" shows "A B C Orth V U" by (meson Orth_def assms orth_at_symmetry) lemma orth_distincts: assumes "A B C Orth U V" shows "A \ B \ B \ C \ A \ C \ U \ V" using Orth_def assms orth_at_distincts by blast lemma col_cop_orth__orth_at: assumes "A B C Orth U V" and "Coplanar A B C X" and "Col U V X" shows "X OrthAt A B C U V" proof - obtain Y where P1: "\ Col A B C \ U \ V \ Coplanar A B C Y \ Col U V Y \ (\ P Q. (Coplanar A B C P \ Col U V Q) \ Per P Y Q)" by (metis OrthAt_def Tarski_neutral_dimensionless.Orth_def Tarski_neutral_dimensionless_axioms assms(1)) then have P2: "X = Y" using assms(2) assms(3) per_distinct_1 by blast { fix P Q assume "Coplanar A B C P \ Col U V Q" then have "Per P X Q" using P1 P2 by auto } thus ?thesis using OrthAt_def Orth_def assms(1) assms(2) assms(3) by auto qed lemma l11_60_aux: assumes "\ Col A B C" and "Cong A P A Q" and "Cong B P B Q" and "Cong C P C Q" and "Coplanar A B C D" shows "Cong D P D Q" proof - obtain M where P1: "Bet P M Q \ Cong P M M Q" by (meson Midpoint_def Tarski_neutral_dimensionless.midpoint_existence Tarski_neutral_dimensionless_axioms) obtain X where P2: " (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 assms(5) Coplanar_def by auto { assume "Col A B X \ Col C D X" then have "Cong D P D Q" by (metis (no_types, lifting) assms(1) assms(2) assms(3) assms(4) l4_17 not_col_distincts not_col_permutation_5) } { assume "Col A C X \ Col B D X" then have "Cong D P D Q" by (metis (no_types, lifting) assms(1) assms(2) assms(3) assms(4) l4_17 not_col_distincts not_col_permutation_5) } { assume "Col A D X \ Col B C X" then have "Cong D P D Q" by (smt assms(1) assms(2) assms(3) assms(4) l4_17 not_col_distincts not_col_permutation_1) } thus ?thesis using P2 \Col A B X \ Col C D X \ Cong D P D Q\ \Col A C X \ Col B D X \ Cong D P D Q\ by blast qed lemma l11_60: assumes "\ Col A B C" and "Per A D P" and "Per B D P" and "Per C D P" and "Coplanar A B C E" shows "Per E D P" by (meson Per_def assms(1) assms(2) assms(3) assms(4) assms(5) l11_60_aux per_double_cong) lemma l11_60_bis: assumes "\ Col A B C" and "D \ P" and "Coplanar A B C D" and "Per A D P" and "Per B D P" and "Per C D P" shows "D OrthAt A B C D P" using assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) l11_60 orth_at_chara by auto lemma l11_61: assumes "A \ A'" and "A \ B" and "A \ C" and "Coplanar A A' B B'" and "Per B A A'" and "Per B' A' A" and "Coplanar A A' C C'" and "Per C A A'" and "Per B A C" shows "Per B' A' C'" proof - have P1: "\ Col C A A'" using assms(1) assms(3) assms(8) per_col_eq by blast obtain C'' where P2: "A A' Perp C'' A' \ A A' OS C C''" using l10_15 using Col_perm P1 col_trivial_2 by blast have P6: "B' \ A" using assms(1) assms(6) per_distinct by blast have P8: "\ Col A' A C''" using P2 not_col_permutation_4 one_side_not_col124 by blast have P9: "Per A' A' B'" by (simp add: l8_2 l8_5) have P10: "Per A A' B'" by (simp add: assms(6) l8_2) { fix B' assume "A A' OS B B' \ Per B' A' A" then have "B A C CongA B' A' C''" using l11_17 by (meson P2 Perp_cases Tarski_neutral_dimensionless.l11_57 Tarski_neutral_dimensionless_axioms assms(5) assms(8) perp_per_1) then have "Per B' A' C''" using assms(9) l11_17 by blast } then have Q1: "\ B'. (A A' OS B B' \ Per B' A' A) \ Per B' A' C''" by simp { fix B' assume P12: "Coplanar A A' B B' \ Per B' A' A \ B' \ A" have "Per B' A' C''" proof cases assume "B' = A'" thus ?thesis by (simp add: Per_perm l8_5) next assume P13: "B' \ A'" have P14: "\ Col B' A' A" using P12 P13 assms(1) l8_9 by auto have P15: "\ Col B A A'" using assms(1) assms(2) assms(5) per_not_col by auto then have Z1: "A A' TS B B' \ A A' OS B B'" using P12 P14 cop__one_or_two_sides not_col_permutation_5 by blast { assume "A A' OS B B'" then have "Per B' A' C''" by (simp add: P12 \\B'a. A A' OS B B'a \ Per B'a A' A \ Per B'a A' C''\) } { assume Q2: "A A' TS B B'" obtain B'' where Z2: "Bet B' A' B'' \ Cong A' B'' A' B'" using segment_construction by blast have "B' \ B''" using P13 Z2 bet_neq12__neq by blast then have Z4: "A' \ B''" using Z2 cong_diff_4 by blast then have "A A' TS B'' B'" by (meson TS_def Z2 Q2 bet__ts invert_two_sides l9_2 not_col_permutation_1) then have Z5: "A A' OS B B''" using Q2 l9_8_1 by auto have "Per B'' A' A" using P12 P13 Z2 bet_col col_per2__per l8_2 l8_5 by blast then have "Per C'' A' B''" using l8_2 Q1 Z5 by blast then have "Per B' A' C''" by (metis Col_def Per_perm Tarski_neutral_dimensionless.l8_3 Tarski_neutral_dimensionless_axioms Z2 Z4) } thus ?thesis using Z1 using \A A' OS B B' \ Per B' A' C''\ by blast qed } then have "\ B'. (Coplanar A A' B B' \ Per B' A' A \ B' \ A) \ Per B' A' C''" by simp then have "Per B' A' C''" using P6 assms(4) assms(6) by blast then have P11: "Per C'' A' B'" using Per_cases by auto have "Coplanar A' A C'' C'" by (meson P1 P2 assms(7) coplanar_trans_1 ncoplanar_perm_6 ncoplanar_perm_8 os__coplanar) thus ?thesis using P8 P9 P10 P11 l8_2 l11_60 by blast qed lemma l11_61_bis: assumes "D OrthAt A B C D P" and "D E Perp E Q" and "Coplanar A B C E" and "Coplanar D E P Q" shows "E OrthAt A B C E Q" proof - have P4: "D \ E" using assms(2) perp_not_eq_1 by auto have P5: "E \ Q" using assms(2) perp_not_eq_2 by auto have "\ D'. (D E Perp D' D \ Coplanar A B C D')" proof - obtain F where T1: "Coplanar A B C F \ \ Col D E F" using P4 ex_ncol_cop by blast obtain D' where T2: "D E Perp D' D \ Coplanar D E F D'" using P4 ex_perp_cop by blast have "Coplanar A B C D'" proof - have T3A: "\ Col A B C" using OrthAt_def assms(1) by auto have T3B: "Coplanar A B C D" using OrthAt_def assms(1) by blast then have T4: "Coplanar D E F A" by (meson T1 T3A assms(3) coplanar_pseudo_trans ncop_distincts) have T5: "Coplanar D E F B" using T1 T3A T3B assms(3) coplanar_pseudo_trans ncop_distincts by blast have "Coplanar D E F C" using T1 T3A T3B assms(3) coplanar_pseudo_trans ncop_distincts by blast thus ?thesis using T1 T2 T4 T5 coplanar_pseudo_trans by blast qed thus ?thesis using T2 by auto qed then obtain D' where R1: "D E Perp D' D \ Coplanar A B C D'" by auto then have R2: "D \ D'" using perp_not_eq_2 by blast { fix M assume R3: "Coplanar A B C M" have "Col D P P" by (simp add: col_trivial_2) then have "Per E D P" using assms(1) assms(3) orth_at_chara by auto then have R4: "Per P D E" using l8_2 by auto have R5: "Per Q E D" using Perp_cases assms(2) perp_per_2 by blast have R6: "Coplanar D E D' M" proof - have S1: "\ Col A B C" using OrthAt_def assms(1) by auto have "Coplanar A B C D" using OrthAt_def assms(1) by auto thus ?thesis using S1 assms(3) R1 R3 coplanar_pseudo_trans by blast qed have R7: "Per D' D E" using Perp_cases R1 perp_per_1 by blast have "Per D' D P" using R1 assms(1) orth_at_chara by blast then have "Per P D D'" using Per_cases by blast then have "Per Q E M" using l11_61 R4 R5 R6 R7 OrthAt_def P4 R2 assms(1) assms(4) by blast then have "Per M E Q" using l8_2 by auto } { fix P0 Q0 assume "Coplanar A B C P0 \ Col E Q Q0" then have "Per P0 E Q0" using P5 \\M. Coplanar A B C M \ Per M E Q\ per_col by blast } thus ?thesis using OrthAt_def P5 assms(1) assms(3) col_trivial_3 by auto qed lemma l11_62_unicity: assumes "Coplanar A B C D" and "Coplanar A B C D'" and "\ E. Coplanar A B C E \ Per E D P" and "\ E. Coplanar A B C E \ Per E D' P" shows "D = D'" by (metis assms(1) assms(2) assms(3) assms(4) l8_8 not_col_distincts per_not_colp) lemma l11_62_unicity_bis: assumes "X OrthAt A B C X U" and "Y OrthAt A B C Y U" shows "X = Y" proof - have P1: "Coplanar A B C X" using assms(1) orth_at_chara by blast have P2: "Coplanar A B C Y" using assms(2) orth_at_chara by blast { fix E assume "Coplanar A B C E" then have "Per E X U" using OrthAt_def assms(1) col_trivial_2 by auto } { fix E assume "Coplanar A B C E" then have "Per E Y U" using assms(2) orth_at_chara by auto } thus ?thesis by (meson P1 P2 \\E. Coplanar A B C E \ Per E X U\ l8_2 l8_7) qed lemma orth_at2__eq: assumes "X OrthAt A B C U V" and "Y OrthAt A B C U V" shows "X = Y" proof - have P1: "Coplanar A B C X" using assms(1) by (simp add: OrthAt_def) have P2: "Coplanar A B C Y" using OrthAt_def assms(2) by auto { fix E assume "Coplanar A B C E" then have "Per E X U" using OrthAt_def assms(1) col_trivial_3 by auto } { fix E assume "Coplanar A B C E" then have "Per E Y U" using OrthAt_def assms(2) col_trivial_3 by auto } thus ?thesis by (meson P1 P2 Per_perm \\E. Coplanar A B C E \ Per E X U\ l8_7) qed lemma col_cop_orth_at__eq: assumes "X OrthAt A B C U V" and "Coplanar A B C Y" and "Col U V Y" shows "X = Y" proof - have "Y OrthAt A B C U V" using Orth_def assms(1) assms(2) assms(3) col_cop_orth__orth_at by blast thus ?thesis using assms(1) orth_at2__eq by auto qed lemma orth_at__ncop1: assumes "U \ X" and "X OrthAt A B C U V" shows "\ Coplanar A B C U" using assms(1) assms(2) col_cop_orth_at__eq not_col_distincts by blast lemma orth_at__ncop2: assumes "V \ X" and "X OrthAt A B C U V" shows "\ Coplanar A B C V" using assms(1) assms(2) col_cop_orth_at__eq not_col_distincts by blast lemma orth_at__ncop: assumes "X OrthAt A B C X P" shows "\ Coplanar A B C P" by (metis assms orth_at__ncop2 orth_at_distincts) lemma l11_62_existence: "\ D. (Coplanar A B C D \ (\ E. (Coplanar A B C E \ Per E D P)))" proof cases assume "Coplanar A B C P" thus ?thesis using l8_5 by auto next assume P1: "\ Coplanar A B C P" then have P2: "\ Col A B C" using ncop__ncol by auto have "\ Col A B P" using P1 ncop__ncols by auto then obtain D0 where P4: "Col A B D0 \ A B Perp P D0" using l8_18_existence by blast have P5: "Coplanar A B C D0" using P4 ncop__ncols by auto have "A \ B" using P2 not_col_distincts by auto then obtain D1 where P10: "A B Perp D1 D0 \ Coplanar A B C D1" using ex_perp_cop by blast have P11: "\ Col A B D1" using P10 P4 perp_not_col2 by blast { fix D assume "Col D0 D1 D" then have "Coplanar A B C D" by (metis P10 P5 col_cop2__cop perp_not_eq_2) } obtain A0 where P11: "A \ A0 \ B \ A0 \ D0 \ A0 \ Col A B A0" using P4 diff_col_ex3 by blast have P12: "Coplanar A B C A0" using P11 ncop__ncols by blast have P13: "Per P D0 A0" using l8_16_1 P11 P4 by blast show ?thesis proof cases assume Z1: "Per P D0 D1" { fix E assume R1: "Coplanar A B C E" have R2: "\ Col A0 D1 D0" by (metis P10 P11 P4 col_permutation_5 colx perp_not_col2) have R3: "Per A0 D0 P" by (simp add: P13 l8_2) have R4: "Per D1 D0 P" by (simp add: Z1 l8_2) have R5: "Per D0 D0 P" by (simp add: l8_2 l8_5) have "Coplanar A0 D1 D0 E" using R1 P2 P12 P10 P5 coplanar_pseudo_trans by blast then have "Per E D0 P" using l11_60 R2 R3 R4 R5 by blast } thus ?thesis using P5 by auto next assume S1: "\ Per P D0 D1" { assume S2: "Col D0 D1 P" have "\ D. Col D0 D1 D \ Coplanar A B C D" by (simp add: \\Da. Col D0 D1 Da \ Coplanar A B C Da\) then have "False" using P1 S2 by blast } then have S2A: "\ Col D0 D1 P" by blast then obtain D where S3: "Col D0 D1 D \ D0 D1 Perp P D" using l8_18_existence by blast have S4: "Coplanar A B C D" by (simp add: S3 \\Da. Col D0 D1 Da \ Coplanar A B C Da\) { fix E assume S5: "Coplanar A B C E" have S6: "D \ D0" using S1 S3 l8_2 perp_per_1 by blast have S7: "Per D0 D P" by (metis Perp_cases S3 S6 perp_col perp_per_1) have S8: "Per D D0 A0" proof - have V4: "D0 \ D1" using P10 perp_not_eq_2 by blast have V6: "Per A0 D0 D1" using P10 P11 P4 l8_16_1 l8_2 by blast thus ?thesis using S3 V4 V6 l8_2 per_col by blast qed have S9: "Per A0 D P" proof - obtain A0' where W1: "D Midpoint A0 A0'" using symmetric_point_construction by auto obtain D0' where W2: "D Midpoint D0 D0'" using symmetric_point_construction by auto have "Cong P A0 P A0'" proof - have V3: "Cong P D0 P D0'" using S7 W2 l8_2 per_double_cong by blast have V4: "Cong D0 A0 D0' A0'" using W1 W2 cong_4321 l7_13 by blast have "Per P D0' A0'" proof - obtain P' where V5: "D Midpoint P P'" using symmetric_point_construction by blast have "Per P' D0 A0" proof - have "\ Col P D D0" by (metis S2A S3 S6 col2__eq col_permutation_1) thus ?thesis by (metis (full_types) P13 S3 S8 V5 S2A col_per2__per midpoint_col) qed thus ?thesis using midpoint_preserves_per V5 Mid_cases W1 W2 by blast qed thus ?thesis using l10_12 P13 V3 V4 by blast qed thus ?thesis using Per_def Per_perm W1 by blast qed have S13: "Per D D P" using Per_perm l8_5 by blast have S14: "\ Col D0 A0 D" using P11 S7 S9 per_not_col Col_perm S6 S8 by blast have "Coplanar A B C D" using S4 by auto then have "Coplanar D0 A0 D E" using P12 P2 P5 S5 coplanar_pseudo_trans by blast then have "Per E D P" using S13 S14 S7 S9 l11_60 by blast } thus ?thesis using S4 by blast qed qed lemma l11_62_existence_bis: assumes "\ Coplanar A B C P" shows "\ X. X OrthAt A B C X P" proof - obtain X where P1: "Coplanar A B C X \ (\ E. Coplanar A B C E \ Per E X P)" using l11_62_existence by blast then have P2: "X \ P" using assms by auto have P3: "\ Col A B C" using assms ncop__ncol by auto thus ?thesis using P1 P2 P3 orth_at_chara by auto qed lemma l11_63_aux: assumes "Coplanar A B C D" and "D \ E" and "E OrthAt A B C E P" shows "\ Q. (D E OS P Q \ A B C Orth D Q)" proof - have P1: "\ Col A B C" using OrthAt_def assms(3) by blast have P2: "E \ P" using OrthAt_def assms(3) by blast have P3: "Coplanar A B C E" using OrthAt_def assms(3) by blast have P4: "\ P0 Q. (Coplanar A B C P0 \ Col E P Q) \ Per P0 E Q" using OrthAt_def assms(3) by blast have P5: "\ Coplanar A B C P" using assms(3) orth_at__ncop by auto have P6: "Col D E D" by (simp add: col_trivial_3) have "\ Col D E P" using P3 P5 assms(1) assms(2) col_cop2__cop by blast then obtain Q where P6: "D E Perp Q D \ D E OS P Q" using P6 l10_15 by blast have "A B C Orth D Q" proof - obtain F where P7: "Coplanar A B C F \ \ Col D E F" using assms(2) ex_ncol_cop by blast obtain D' where P8: "D E Perp D' D \ Coplanar D E F D'" using assms(2) ex_perp_cop by presburger have P9: "\ Col D' D E" using P8 col_permutation_1 perp_not_col by blast have P10: "Coplanar D E F A" by (meson P1 P3 P7 assms(1) coplanar_pseudo_trans ncop_distincts) have P11: "Coplanar D E F B" by (meson P1 P3 P7 assms(1) coplanar_pseudo_trans ncop_distincts) have P12: "Coplanar D E F C" by (meson P1 P3 P7 assms(1) coplanar_pseudo_trans ncop_distincts) then have "D OrthAt A B C D Q" proof - have "Per D' D Q" proof - obtain E' where Y1: "D E Perp E' E \ Coplanar D E F E'" using assms(2) ex_perp_cop by blast have Y2: "E \ E'" using Y1 perp_distinct by auto have Y5: "Coplanar E D E' D'" by (meson P7 P8 Y1 coplanar_perm_12 coplanar_perm_7 coplanar_trans_1 not_col_permutation_2) have Y6: "Per E' E D" by (simp add: Perp_perm Tarski_neutral_dimensionless.perp_per_2 Tarski_neutral_dimensionless_axioms Y1) have Y7: "Per D' D E" using P8 col_trivial_2 col_trivial_3 l8_16_1 by blast have Y8: "Coplanar E D P Q" using P6 ncoplanar_perm_6 os__coplanar by blast have Y9: "Per P E D" using P4 using assms(1) assms(3) l8_2 orth_at_chara by blast have Y10: "Coplanar A B C E'" using P10 P11 P12 P7 Y1 coplanar_pseudo_trans by blast then have Y11: "Per E' E P" using P4 Y10 col_trivial_2 by auto have "E \ D" using assms(2) by blast thus ?thesis using l11_61 Y2 assms(2) P2 Y5 Y6 Y7 Y8 Y9 Y10 Y11 by blast qed then have X1: "D OrthAt D' D E D Q" using l11_60_bis by (metis OS_def P6 P9 Per_perm TS_def Tarski_neutral_dimensionless.l8_5 Tarski_neutral_dimensionless_axioms col_trivial_3 invert_one_side ncop_distincts perp_per_1) have X3: "Coplanar D' D E A" using P10 P7 P8 coplanar_perm_14 coplanar_trans_1 not_col_permutation_3 by blast have X4: "Coplanar D' D E B" using P11 P7 P8 coplanar_perm_14 coplanar_trans_1 not_col_permutation_3 by blast have "Coplanar D' D E C" using P12 P7 P8 coplanar_perm_14 coplanar_trans_1 not_col_permutation_3 by blast thus ?thesis using X1 P1 X3 X4 cop3_orth_at__orth_at by blast qed thus ?thesis using Orth_def by blast qed thus ?thesis using P6 by blast qed lemma l11_63_existence: assumes "Coplanar A B C D" and "\ Coplanar A B C P" shows "\ Q. A B C Orth D Q" using Orth_def assms(1) assms(2) l11_62_existence_bis l11_63_aux by fastforce lemma l8_21_3: assumes "Coplanar A B C D" and "\ Coplanar A B C X" shows "\ P T. (A B C Orth D P \ Coplanar A B C T \ Bet X T P)" proof - obtain E where P1: "E OrthAt A B C E X" using assms(2) l11_62_existence_bis by blast thus ?thesis proof cases assume P2: "D = E" obtain Y where P3: "Bet X D Y \ Cong D Y D X" using segment_construction by blast have P4: "D \ X" using assms(1) assms(2) by auto have P5: "A B C Orth D X" using Orth_def P1 P2 by auto have P6: "D \ Y" using P3 P4 cong_reverse_identity by blast have "Col D X Y" using Col_def Col_perm P3 by blast then have "A B C Orth D Y" using P5 P6 col_orth__orth by auto thus ?thesis using P3 assms(1) by blast next assume K1: "D \ E" then obtain P' where P7: "D E OS X P' \ A B C Orth D P'" using P1 assms(1) l11_63_aux by blast have P8: "\ Col A B C" using assms(2) ncop__ncol by auto have P9: "E \ X" using P7 os_distincts by auto have P10: "\ P Q. (Coplanar A B C P \ Col E X Q) \ Per P E Q" using OrthAt_def P1 by auto have P11: "D OrthAt A B C D P'" by (simp add: P7 assms(1) col_cop_orth__orth_at col_trivial_3) have P12: "D \ P'" using P7 os_distincts by auto have P13: "\ Coplanar A B C P'" using P11 orth_at__ncop by auto have P14: "\ P Q. (Coplanar A B C P \ Col D P' Q) \ Per P D Q" using OrthAt_def P11 by auto obtain P where P15: "Bet P' D P \ Cong D P D P'" using segment_construction by blast have P16: "D E TS X P" proof - have P16A: "D E OS P' X" using P7 one_side_symmetry by blast then have "D E TS P' P" by (metis P12 P15 Tarski_neutral_dimensionless.bet__ts Tarski_neutral_dimensionless_axioms cong_diff_3 one_side_not_col123) thus ?thesis using l9_8_2 P16A by blast qed obtain T where P17: "Col T D E \ Bet X T P" using P16 TS_def by blast have P18: "D \ P" using P16 ts_distincts by blast have "Col D P' P" using Col_def Col_perm P15 by blast then have "A B C Orth D P" using P18 P7 col_orth__orth by blast thus ?thesis using col_cop2__cop by (meson P1 P17 Tarski_neutral_dimensionless.orth_at_chara Tarski_neutral_dimensionless_axioms K1 assms(1) col_permutation_1) qed qed lemma mid2_orth_at2__cong: assumes "X OrthAt A B C X P" and "Y OrthAt A B C Y Q" and "X Midpoint P P'" and "Y Midpoint Q Q'" shows "Cong P Q P' Q'" proof - have Q1: "\ Col A B C" using assms(1) col__coplanar orth_at__ncop by blast have Q2: "X \ P" using assms(1) orth_at_distincts by auto have Q3: "Coplanar A B C X" using OrthAt_def assms(1) by auto have Q4: "\ P0 Q. (Coplanar A B C P0 \ Col X P Q) \ Per P0 X Q" using OrthAt_def assms(1) by blast have Q5: "Y \ P" by (metis assms(1) assms(2) orth_at__ncop2 orth_at_chara) have Q6: "Coplanar A B C Y" using OrthAt_def assms(2) by blast have Q7: "\ P Q0. (Coplanar A B C P \ Col Y Q Q0) \ Per P Y Q0" using OrthAt_def assms(2) by blast obtain Z where P1: "Z Midpoint X Y" using midpoint_existence by auto obtain R where P2: "Z Midpoint P R" using symmetric_point_construction by auto obtain R' where P3: "Z Midpoint P' R'" using symmetric_point_construction by auto have T1: "Coplanar A B C Z" using P1 Q3 Q6 bet_cop2__cop midpoint_bet by blast then have "Per Z X P" using Q4 assms(1) orth_at_chara by auto then have T2: "Cong Z P Z P'" using assms(3) per_double_cong by blast have T3: "Cong R Z R' Z" by (metis Cong_perm Midpoint_def P2 P3 T2 cong_transitivity) have T4: "Cong R Q R' Q'" by (meson P1 P2 P3 assms(3) assms(4) l7_13 not_cong_4321 symmetry_preserves_midpoint) have "Per Z Y Q" using Q7 T1 assms(2) orth_at_chara by auto then have T5: "Cong Z Q Z Q'" using assms(4) per_double_cong by auto have "R \ Z" by (metis P2 P3 Q2 T3 assms(3) cong_diff_3 l7_17_bis midpoint_not_midpoint) thus ?thesis using P2 P3 T2 T3 T4 T5 five_segment l7_2 midpoint_bet by blast qed lemma orth_at2_tsp__ts: assumes "P \ Q" and "P OrthAt A B C P X" and "Q OrthAt A B C Q Y" and "A B C TSP X Y" shows "P Q TS X Y" proof - obtain T where P0: "Coplanar A B C T \ Bet X T Y" using TSP_def assms(4) by auto have P1: "\ Col A B C" using assms(4) ncop__ncol tsp__ncop1 by blast have P2: "P \ X " using assms(2) orth_at_distincts by auto have P3: "Coplanar A B C P" using OrthAt_def assms(2) by blast have P4: "\ D. Coplanar A B C D \ Per D P X" using assms(2) orth_at_chara by blast have P5: "Q \ Y" using assms(3) orth_at_distincts by auto have P6: "Coplanar A B C Q" using OrthAt_def assms(3) by blast have P7: "\ D. Coplanar A B C D \ Per D Q Y" using assms(3) orth_at_chara by blast have P8: "\ Col X P Q" using P3 P6 assms(1) assms(4) col_cop2__cop not_col_permutation_2 tsp__ncop1 by blast have P9: "\ Col Y P Q" using P3 P6 assms(1) assms(4) col_cop2__cop not_col_permutation_2 tsp__ncop2 by blast have "Col T P Q" proof - obtain X' where Q1: "P Midpoint X X'" using symmetric_point_construction by auto obtain Y' where Q2: "Q Midpoint Y Y'" using symmetric_point_construction by auto have "Per T P X" using P0 P4 by auto then have Q3: "Cong T X T X'" using Q1 per_double_cong by auto have "Per T Q Y" using P0 P7 by auto then have Q4: "Cong T Y T Y'" using Q2 per_double_cong by auto have "Cong X Y X' Y'" using P1 Q1 Q2 assms(2) assms(3) mid2_orth_at2__cong by blast then have "X T Y Cong3 X' T Y'" using Cong3_def Q3 Q4 not_cong_2143 by blast then have "Bet X' T Y'" using l4_6 P0 by blast thus ?thesis using Q1 Q2 Q3 Q4 Col_def P0 between_symmetry l7_22 by blast qed thus ?thesis using P0 P8 P9 TS_def by blast qed lemma orth_dec: shows "A B C Orth U V \ \ A B C Orth U V" by auto lemma orth_at_dec: shows "X OrthAt A B C U V \ \ X OrthAt A B C U V" by auto lemma tsp_dec: shows "A B C TSP X Y \ \ A B C TSP X Y" by auto lemma osp_dec: shows "A B C OSP X Y \ \ A B C OSP X Y" by auto lemma ts2__inangle: assumes "A C TS B P" and "B P TS A C" shows "P InAngle A B C" by (metis InAngle_def assms(1) assms(2) bet_out ts2__ex_bet2 ts_distincts) lemma os_ts__inangle: assumes "B P TS A C" and "B A OS C P" shows "P InAngle A B C" proof - have P1: "\ Col A B P" using TS_def assms(1) by auto have P2: "\ Col B A C" using assms(2) col123__nos by blast obtain P' where P3: "B Midpoint P P'" using symmetric_point_construction by blast then have P4: "\ Col B A P'" by (metis assms(2) col_one_side col_permutation_5 midpoint_col midpoint_distinct_2 one_side_not_col124) have P5: "(B \ P' \ B P TS A C \ Bet P B P') \ (P InAngle A B C \ P' InAngle A B C)" using two_sides_in_angle by auto then have P6: "P InAngle A B C \ P' InAngle A B C" using P3 P4 assms(1) midpoint_bet not_col_distincts by blast { assume "P' InAngle A B C" then have P7: "A B OS P' C" using Col_cases P2 P4 in_angle_one_side by blast then have P8: "\ A B TS P' C" using l9_9 by auto have "B A TS P P'" using P1 P3 P4 bet__ts midpoint_bet not_col_distincts not_col_permutation_4 by auto then have "B A TS C P'" using P7 assms(2) invert_one_side l9_2 l9_8_2 l9_9 by blast then have "B A TS P' C" using l9_2 by blast then have "A B TS P' C" by (simp add: invert_two_sides) then have "P InAngle A B C" using P8 by auto } thus ?thesis using P6 by blast qed lemma os2__inangle: assumes "B A OS C P" and "B C OS A P" shows "P InAngle A B C" using assms(1) assms(2) col124__nos l9_9_bis os_ts__inangle two_sides_cases by blast lemma acute_conga__acute: assumes "Acute A B C" and "A B C CongA D E F" shows "Acute D E F" proof - have "D E F LeA A B C" by (simp add: assms(2) conga__lea456123) thus ?thesis using acute_lea_acute assms(1) by blast qed lemma acute_out2__acute: assumes "B Out A' A" and "B Out C' C" and "Acute A B C" shows "Acute A' B C'" by (meson Tarski_neutral_dimensionless.out2__conga Tarski_neutral_dimensionless_axioms acute_conga__acute assms(1) assms(2) assms(3)) lemma conga_obtuse__obtuse: assumes "Obtuse A B C" and "A B C CongA D E F" shows "Obtuse D E F" using assms(1) assms(2) conga__lea lea_obtuse_obtuse by blast lemma obtuse_out2__obtuse: assumes "B Out A' A" and "B Out C' C" and "Obtuse A B C" shows "Obtuse A' B C'" by (meson Tarski_neutral_dimensionless.out2__conga Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) conga_obtuse__obtuse) lemma bet_lea__bet: assumes "Bet A B C" and "A B C LeA D E F" shows "Bet D E F" proof - have "A B C CongA D E F" by (metis assms(1) assms(2) l11_31_2 lea_asym lea_distincts) thus ?thesis using assms(1) bet_conga__bet by blast qed lemma out_lea__out: assumes "E Out D F" and "A B C LeA D E F" shows "B Out A C" proof - have "D E F CongA A B C" using Tarski_neutral_dimensionless.l11_31_1 Tarski_neutral_dimensionless.lea_asym Tarski_neutral_dimensionless.lea_distincts Tarski_neutral_dimensionless_axioms assms(1) assms(2) by fastforce thus ?thesis using assms(1) out_conga_out by blast qed lemma bet2_lta__lta: assumes "A B C LtA D E F" and "Bet A B A'" and "A' \ B" and "Bet D E D'" and "D' \ E" shows "D' E F LtA A' B C" proof - have P1: "D' E F LeA A' B C" by (metis Bet_cases assms(1) assms(2) assms(3) assms(4) assms(5) l11_36_aux2 lea_distincts lta__lea) have "\ D' E F CongA A' B C" by (metis assms(1) assms(2) assms(4) between_symmetry conga_sym l11_13 lta_distincts not_lta_and_conga) thus ?thesis by (simp add: LtA_def P1) qed lemma lea123456_lta__lta: assumes "A B C LeA D E F" and "D E F LtA G H I" shows "A B C LtA G H I" proof - have "\ G H I LeA F E D" by (metis (no_types) Tarski_neutral_dimensionless.lea__nlta Tarski_neutral_dimensionless.lta_left_comm Tarski_neutral_dimensionless_axioms assms(2)) then have "\ A B C CongA G H I" by (metis Tarski_neutral_dimensionless.lta_distincts Tarski_neutral_dimensionless_axioms assms(1) assms(2) conga_pseudo_refl l11_30) thus ?thesis by (meson LtA_def Tarski_neutral_dimensionless.lea_trans Tarski_neutral_dimensionless_axioms assms(1) assms(2)) qed lemma lea456789_lta__lta: assumes "A B C LtA D E F" and "D E F LeA G H I" shows "A B C LtA G H I" by (meson LtA_def assms(1) assms(2) conga__lea456123 lea_trans lta__nlea) lemma acute_per__lta: assumes "Acute A B C" and "D \ E" and "E \ F" and "Per D E F" shows "A B C LtA D E F" proof - obtain G H I where P1: "Per G H I \ A B C LtA G H I" using Acute_def assms(1) by auto then have "G H I CongA D E F" using assms(2) assms(3) assms(4) l11_16 lta_distincts by blast thus ?thesis by (metis P1 conga_preserves_lta conga_refl lta_distincts) qed lemma obtuse_per__lta: assumes "Obtuse A B C" and "D \ E" and "E \ F" and "Per D E F" shows "D E F LtA A B C" proof - obtain G H I where P1: "Per G H I \ G H I LtA A B C" using Obtuse_def assms(1) by auto then have "G H I CongA D E F" using assms(2) assms(3) assms(4) l11_16 lta_distincts by blast thus ?thesis by (metis P1 Tarski_neutral_dimensionless.l11_51 Tarski_neutral_dimensionless_axioms assms(1) cong_reflexivity conga_preserves_lta obtuse_distincts) qed lemma acute_obtuse__lta: assumes "Acute A B C" and "Obtuse D E F" shows "A B C LtA D E F" by (metis Acute_def assms(1) assms(2) lta_distincts lta_trans obtuse_per__lta) lemma lea_in_angle: assumes "A B P LeA A B C" and "A B OS C P" shows "P InAngle A B C" proof - obtain T where P3: "T InAngle A B C \ A B P CongA A B T" using LeA_def assms(1) by blast thus ?thesis by (metis assms(2) conga_preserves_in_angle conga_refl not_conga_sym one_side_symmetry os_distincts) qed lemma acute_bet__obtuse: assumes "Bet A B A'" and "A' \ B" and "Acute A B C" shows "Obtuse A' B C" proof cases assume P1: "Col A B C" show ?thesis proof cases assume "Bet A B C" thus ?thesis using P1 acute_col__out assms(3) not_bet_and_out by blast next assume "\ Bet A B C" thus ?thesis by (smt P1 assms(1) assms(2) bet__obtuse between_inner_transitivity between_symmetry outer_transitivity_between third_point) qed next assume P2: "\ Col A B C" then obtain D where P3: "A B Perp D B \ A B OS C D" using col_trivial_2 l10_15 by blast { assume P4: "Col C B D" then have "Per A B C" proof - have P5: "B \ D" using P3 perp_not_eq_2 by auto have "Per A B D" using P3 Perp_perm perp_per_2 by blast thus ?thesis using P4 P5 not_col_permutation_2 per_col by blast qed then have "A B C LtA A B C" by (metis Acute_def acute_per__lta assms(3) lta_distincts) then have "False" by (simp add: nlta) } then have P6: "\ Col C B D" by auto have P7: "B A' OS C D" by (metis P3 assms(1) assms(2) bet_col col2_os__os l5_3) have T1: "Per A B D" by (simp add: P3 perp_left_comm perp_per_1) have Q1: "B C TS A' A" using P2 assms(1) assms(2) bet__ts l9_2 not_col_permutation_1 by auto have "A B C LtA A B D" using P2 P6 T1 acute_per__lta assms(3) not_col_distincts by auto then have "A B C LeA A B D" by (simp add: lta__lea) then have "C InAngle A B D" by (simp add: P3 lea_in_angle one_side_symmetry) then have "C InAngle D B A" using l11_24 by blast then have "C B TS D A" by (simp add: P2 P6 in_angle_two_sides not_col_permutation_1 not_col_permutation_4) then have "B C TS D A" using invert_two_sides by blast then have "B C OS A' D" using Q1 l9_8_1 by auto then have T1A: "D InAngle A' B C" by (simp add: P7 os2__inangle) then have "A B D CongA A' B D" by (metis Per_cases T1 Tarski_neutral_dimensionless.conga_comm Tarski_neutral_dimensionless.l11_18_1 Tarski_neutral_dimensionless_axioms acute_distincts assms(1) assms(3) inangle_distincts) then have T2: "A B D LeA A' B C" using LeA_def T1A by auto { assume "A B D CongA A' B C" then have "False" by (metis OS_def P7 T1 TS_def Tarski_neutral_dimensionless.out2__conga Tarski_neutral_dimensionless_axioms \A B C LtA A B D\ \A B D CongA A' B D\ \\thesis. (\D. A B Perp D B \ A B OS C D \ thesis) \ thesis\ col_trivial_2 invert_one_side l11_17 l11_19 not_lta_and_conga out_trivial) } then have "\ A B D CongA A' B C" by auto then have "A B D LtA A' B C" using T2 LtA_def by auto thus ?thesis using T1 Obtuse_def by blast qed lemma bet_obtuse__acute: assumes "Bet A B A'" and "A' \ B" and "Obtuse A B C" shows "Acute A' B C" proof cases assume P1: "Col A B C" thus ?thesis proof cases assume "Bet A B C" then have "B Out A' C" by (smt Out_def assms(1) assms(2) assms(3) l5_2 obtuse_distincts) thus ?thesis by (simp add: out__acute) next assume "\ Bet A B C" thus ?thesis using P1 assms(3) col_obtuse__bet by blast qed next assume P2: "\ Col A B C" then obtain D where P3: "A B Perp D B \ A B OS C D" using col_trivial_2 l10_15 by blast { assume P3A: "Col C B D" have P3B: "B \ D" using P3 perp_not_eq_2 by blast have P3C: "Per A B D" using P3 Perp_perm perp_per_2 by blast then have "Per A B C" using P3A P3B not_col_permutation_2 per_col by blast then have "A B C LtA A B C" using P2 assms(3) not_col_distincts obtuse_per__lta by auto then have "False" by (simp add: nlta) } then have P4: "\ Col C B D" by auto have "Col B A A'" using Col_def Col_perm assms(1) by blast then have P5: "B A' OS C D" using P3 assms(2) invert_one_side col2_os__os col_trivial_3 by blast have P7: "Per A' B D" by (meson Col_def P3 Tarski_neutral_dimensionless.Per_perm Tarski_neutral_dimensionless_axioms assms(1) col_trivial_2 l8_16_1) have "A' B C LtA A' B D" proof - have P8: "A' B C LeA A' B D" proof - have P9: "C InAngle A' B D" proof - have P10: "B A' OS D C" by (simp add: P5 one_side_symmetry) have "B D OS A' C" proof - have P6: "\ Col A B D" using P3 col124__nos by auto then have P11: "B D TS A' A" using Col_perm P5 assms(1) bet__ts l9_2 os_distincts by blast have "A B D LtA A B C" proof - have P11A: "A \ B" using P2 col_trivial_1 by auto have P11B: "B \ D" using P4 col_trivial_2 by blast have "Per A B D" using P3 Perp_cases perp_per_2 by blast thus ?thesis by (simp add: P11A P11B Tarski_neutral_dimensionless.obtuse_per__lta Tarski_neutral_dimensionless_axioms assms(3)) qed then have "A B D LeA A B C" by (simp add: lta__lea) then have "D InAngle A B C" by (simp add: P3 lea_in_angle) then have "D InAngle C B A" using l11_24 by blast then have "D B TS C A" by (simp add: P4 P6 in_angle_two_sides not_col_permutation_4) then have "B D TS C A" by (simp add: invert_two_sides) thus ?thesis using OS_def P11 by blast qed thus ?thesis by (simp add: P10 os2__inangle) qed have "A' B C CongA A' B C" using assms(2) assms(3) conga_refl obtuse_distincts by blast thus ?thesis by (simp add: P9 inangle__lea) qed { assume "A' B C CongA A' B D" then have "B Out C D" using P5 conga_os__out invert_one_side by blast then have "False" using P4 not_col_permutation_4 out_col by blast } then have "\ A' B C CongA A' B D" by auto thus ?thesis by (simp add: LtA_def P8) qed thus ?thesis using Acute_def P7 by blast qed lemma inangle_dec: "P InAngle A B C \ \ P InAngle A B C" by blast lemma lea_dec: "A B C LeA D E F \ \ A B C LeA D E F" by blast lemma lta_dec: "A B C LtA D E F \ \ A B C LtA D E F" by blast lemma lea_total: assumes "A \ B" and "B \ C" and "D \ E" and "E \ F" shows "A B C LeA D E F \ D E F LeA A B C" proof cases assume P1: "Col A B C" show ?thesis proof cases assume "B Out A C" thus ?thesis using assms(3) assms(4) l11_31_1 by auto next assume "\ B Out A C" thus ?thesis by (metis P1 assms(1) assms(2) assms(3) assms(4) l11_31_2 or_bet_out) qed next assume P2: "\ Col A B C" show ?thesis proof cases assume P3: "Col D E F" show ?thesis proof cases assume "E Out D F" thus ?thesis using assms(1) assms(2) l11_31_1 by auto next assume "\ E Out D F" thus ?thesis by (metis P3 assms(1) assms(2) assms(3) assms(4) l11_31_2 l6_4_2) qed next assume P4: "\ Col D E F" show ?thesis proof cases assume "A B C LeA D E F" thus ?thesis by simp next assume P5: "\ A B C LeA D E F" obtain P where P6: "D E F CongA A B P \ A B OS P C" using P2 P4 angle_construction_1 by blast then have P7: "B A OS C P" using invert_one_side one_side_symmetry by blast have "B C OS A P" proof - { assume "Col P B C" then have P7B: "B Out C P" using Col_cases P7 col_one_side_out by blast have "A B C CongA D E F" proof - have P7C: "A B P CongA D E F" by (simp add: P6 conga_sym) have P7D: "B Out A A" by (simp add: assms(1) out_trivial) have P7E: "E Out D D" by (simp add: assms(3) out_trivial) have "E Out F F" using assms(4) out_trivial by auto thus ?thesis using P7B P7C P7D P7E l11_10 by blast qed then have "A B C LeA D E F" by (simp add: conga__lea) then have "False" by (simp add: P5) } then have P8: "\ Col P B C" by auto { assume T0: "B C TS A P" have "A B C CongA D E F" proof - have T1: "A B C LeA A B P" proof - have T1A: "C InAngle A B P" by (simp add: P7 T0 one_side_symmetry os_ts__inangle) have "A B C CongA A B C" using assms(1) assms(2) conga_refl by auto thus ?thesis by (simp add: T1A inangle__lea) qed have T2: "A B C CongA A B C" using assms(1) assms(2) conga_refl by auto have "A B P CongA D E F" by (simp add: P6 conga_sym) thus ?thesis using P5 T1 T2 l11_30 by blast qed then have "A B C LeA D E F" by (simp add: conga__lea) then have "False" by (simp add: P5) } then have "\ B C TS A P" by auto thus ?thesis using Col_perm P7 P8 one_side_symmetry os_ts1324__os two_sides_cases by blast qed then have "P InAngle A B C" using P7 os2__inangle by blast then have "D E F LeA A B C" using P6 LeA_def by blast thus ?thesis by simp qed qed qed lemma or_lta2_conga: assumes "A \ B" and "C \ B" and "D \ E" and "F \ E" shows "A B C LtA D E F \ D E F LtA A B C \ A B C CongA D E F" proof - have P1: "A B C LeA D E F \ D E F LeA A B C" using assms(1) assms(2) assms(3) assms(4) lea_total by auto { assume "A B C LeA D E F" then have "A B C LtA D E F \ D E F LtA A B C \ A B C CongA D E F" using LtA_def by blast } { assume "D E F LeA A B C" then have "A B C LtA D E F \ D E F LtA A B C \ A B C CongA D E F" using LtA_def conga_sym by blast } thus ?thesis using P1 \A B C LeA D E F \ A B C LtA D E F \ D E F LtA A B C \ A B C CongA D E F\ by blast qed lemma angle_partition: assumes "A \ B" and "B \ C" shows "Acute A B C \ Per A B C \ Obtuse A B C" proof - obtain A' B' D' where P1: "\ (Bet A' B' D' \ Bet B' D' A' \ Bet D' A' B')" using lower_dim by auto then have "\ Col A' B' D'" by (simp add: Col_def) obtain C' where P3: "A' B' Perp C' B'" by (metis P1 Perp_perm between_trivial2 perp_exists) then have P4: "A B C LtA A' B' C' \ A' B' C' LtA A B C \ A B C CongA A' B' C'" by (metis P1 assms(1) assms(2) between_trivial2 or_lta2_conga perp_not_eq_2) { assume "A B C LtA A' B' C'" then have "Acute A B C \ Per A B C \ Obtuse A B C" using Acute_def P3 Perp_cases perp_per_2 by blast } { assume "A' B' C' LtA A B C" then have "Acute A B C \ Per A B C \ Obtuse A B C" using Obtuse_def P3 Perp_cases perp_per_2 by blast } { assume "A B C CongA A' B' C'" then have "Acute A B C \ Per A B C \ Obtuse A B C" by (metis P3 Perp_cases Tarski_neutral_dimensionless.conga_sym Tarski_neutral_dimensionless.l11_17 Tarski_neutral_dimensionless_axioms perp_per_2) } thus ?thesis using P4 \A B C LtA A' B' C' \ Acute A B C \ Per A B C \ Obtuse A B C\ \A' B' C' LtA A B C \ Acute A B C \ Per A B C \ Obtuse A B C\ by auto qed lemma acute_chara_1: assumes "Bet A B A'" and "B \ A'" and "Acute A B C" shows "A B C LtA A' B C" proof - have "Obtuse A' B C" using acute_bet__obtuse assms(1) assms(2) assms(3) by blast thus ?thesis by (simp add: acute_obtuse__lta assms(3)) qed lemma acute_chara_2: assumes "Bet A B A'" and "A B C LtA A' B C" shows "Acute A B C" by (metis Tarski_neutral_dimensionless.Acute_def Tarski_neutral_dimensionless_axioms acute_chara_1 angle_partition assms(1) assms(2) bet_obtuse__acute between_symmetry lta_distincts not_and_lta) lemma acute_chara: assumes "Bet A B A'" and "B \ A'" shows "Acute A B C \ A B C LtA A' B C" using acute_chara_1 acute_chara_2 assms(1) assms(2) by blast lemma obtuse_chara: assumes "Bet A B A'" and "B \ A'" shows "Obtuse A B C \ A' B C LtA A B C" by (metis Tarski_neutral_dimensionless.Obtuse_def Tarski_neutral_dimensionless_axioms acute_bet__obtuse acute_chara assms(1) assms(2) bet_obtuse__acute between_symmetry lta_distincts) lemma conga__acute: assumes "A B C CongA A C B" shows "Acute A B C" by (metis acute_sym angle_partition assms conga_distinct conga_obtuse__obtuse l11_17 l11_43) lemma cong__acute: assumes "A \ B" and "B \ C" and "Cong A B A C" shows "Acute A B C" using angle_partition assms(1) assms(2) assms(3) cong__nlt l11_46 lt_left_comm by blast lemma nlta__lea: assumes "\ A B C LtA D E F" and "A \ B" and "B \ C" and "D \ E" and "E \ F" shows "D E F LeA A B C" by (metis LtA_def assms(1) assms(2) assms(3) assms(4) assms(5) conga__lea456123 or_lta2_conga) lemma nlea__lta: assumes "\ A B C LeA D E F" and "A \ B" and "B \ C" and "D \ E" and "E \ F" shows "D E F LtA A B C" using assms(1) assms(2) assms(3) assms(4) assms(5) nlta__lea by blast lemma triangle_strict_inequality: assumes "Bet A B D" and "Cong B C B D" and "\ Bet A B C" shows "A C Lt A D" proof cases assume P1: "Col A B C" then have P2: "B Out A C" using assms(3) not_out_bet by auto { assume "Bet B A C" then have P3: "A C Le A D" by (meson assms(1) assms(2) cong__le l5_12_a le_transitivity) have "\ Cong A C A D" by (metis Out_def P1 P2 assms(1) assms(2) assms(3) l4_18) then have "A C Lt A D" by (simp add: Lt_def P3) } { assume "Bet A C B" then have P5: "Bet A C D" using assms(1) between_exchange4 by blast then have P6: "A C Le A D" by (simp add: bet__le1213) have "\ Cong A C A D" using P5 assms(1) assms(3) between_cong by blast then have "A C Lt A D" by (simp add: Lt_def P6) } thus ?thesis using P1 \Bet B A C \ A C Lt A D\ assms(3) between_symmetry third_point by blast next assume T1: "\ Col A B C" have T2: "A \ D" using T1 assms(1) between_identity col_trivial_1 by auto have T3: "\ Col A C D" by (metis Col_def T1 T2 assms(1) col_transitivity_2) have T4: "Bet D B A" using Bet_perm assms(1) by blast have T5: "C D A CongA D C B" proof - have T6: "C D B CongA D C B" by (metis assms(1) assms(2) assms(3) between_trivial conga_comm l11_44_1_a not_conga_sym) have T7: "D Out C C" using T3 not_col_distincts out_trivial by blast have T8: "D Out A B" by (metis assms(1) assms(2) assms(3) bet_out_1 cong_diff l6_6) have T9: "C Out D D" using T7 out_trivial by force have "C Out B B" using T1 not_col_distincts out_trivial by auto thus ?thesis using T6 T7 T8 T9 l11_10 by blast qed have "A D C LtA A C D" proof - have "B InAngle D C A" by (metis InAngle_def T1 T3 T4 not_col_distincts out_trivial) then have "C D A LeA D C A" using T5 LeA_def by auto then have T10: "A D C LeA A C D" by (simp add: lea_comm) have "\ A D C CongA A C D" by (metis Col_perm T3 assms(1) assms(2) assms(3) bet_col l11_44_1_b l4_18 not_bet_distincts not_cong_3412) thus ?thesis using LtA_def T10 by blast qed thus ?thesis by (simp add: l11_44_2_b) qed lemma triangle_inequality: assumes "Bet A B D" and "Cong B C B D" shows "A C Le A D" proof cases assume "Bet A B C" thus ?thesis by (metis assms(1) assms(2) between_cong_3 cong__le le_reflexivity) next assume "\ Bet A B C" then have "A C Lt A D" using assms(1) assms(2) triangle_strict_inequality by auto thus ?thesis by (simp add: Lt_def) qed lemma triangle_strict_inequality_2: assumes "Bet A' B' C'" and "Cong A B A' B'" and "Cong B C B' C'" and "\ Bet A B C" shows "A C Lt A' C'" proof - obtain D where P1: "Bet A B D \ Cong B D B C" using segment_construction by blast then have P2: "A C Lt A D" using P1 assms(4) cong_symmetry triangle_strict_inequality by blast have "Cong A D A' C'" using P1 assms(1) assms(2) assms(3) cong_transitivity l2_11_b by blast thus ?thesis using P2 cong2_lt__lt cong_reflexivity by blast qed lemma triangle_inequality_2: assumes "Bet A' B' C'" and "Cong A B A' B'" and "Cong B C B' C'" shows "A C Le A' C'" proof - obtain D where P1: "Bet A B D \ Cong B D B C" using segment_construction by blast have P2: "A C Le A D" by (meson P1 Tarski_neutral_dimensionless.triangle_inequality Tarski_neutral_dimensionless_axioms not_cong_3412) have "Cong A D A' C'" using P1 assms(1) assms(2) assms(3) cong_transitivity l2_11_b by blast thus ?thesis using P2 cong__le le_transitivity by blast qed lemma triangle_strict_reverse_inequality: assumes "A Out B D" and "Cong A C A D" and "\ A Out B C" shows "B D Lt B C" proof cases assume "Col A B C" thus ?thesis using assms(1) assms(2) assms(3) col_permutation_4 cong_symmetry not_bet_and_out or_bet_out triangle_strict_inequality by blast next assume P1: "\ Col A B C" show ?thesis proof cases assume "B = D" thus ?thesis using P1 lt1123 not_col_distincts by auto next assume P2: "B \ D" then have P3: "\ Col B C D" using P1 assms(1) col_trivial_2 colx not_col_permutation_5 out_col by blast have P4: "\ Col A C D" using P1 assms(1) col2__eq col_permutation_4 out_col out_distinct by blast have P5: "C \ D" using assms(1) assms(3) by auto then have P6: "A C D CongA A D C" by (metis P1 assms(2) col_trivial_3 l11_44_1_a) { assume T1: "Bet A B D" then have T2: "Bet D B A" using Bet_perm by blast have "B C D LtA B D C" proof - have T3: "D C B CongA B C D" by (metis P3 conga_pseudo_refl not_col_distincts) have T3A: "D Out B A" by (simp add: P2 T1 bet_out_1) have T3B: "C Out D D" using P5 out_trivial by auto have T3C: "C Out A A" using P1 not_col_distincts out_trivial by blast have "D Out C C" by (simp add: P5 out_trivial) then have T4: "D C A CongA B D C" using T3A T3B T3C by (meson Tarski_neutral_dimensionless.conga_comm Tarski_neutral_dimensionless.conga_right_comm Tarski_neutral_dimensionless.l11_10 Tarski_neutral_dimensionless_axioms P6) have "D C B LtA D C A" proof - have T4A: "D C B LeA D C A" proof - have T4AA: "B InAngle D C A" using InAngle_def P1 P5 T2 not_col_distincts out_trivial by auto have "D C B CongA D C B" using T3 conga_right_comm by blast thus ?thesis by (simp add: T4AA inangle__lea) qed { assume T5: "D C B CongA D C A" have "D C OS B A" using Col_perm P3 T3A out_one_side by blast then have "C Out B A" using T5 conga_os__out by blast then have "False" using Col_cases P1 out_col by blast } then have "\ D C B CongA D C A" by auto thus ?thesis using LtA_def T4A by blast qed thus ?thesis using T3 T4 conga_preserves_lta by auto qed } { assume Q1: "Bet B D A" obtain E where Q2: "Bet B C E \ Cong B C C E" using Cong_perm segment_construction by blast have "A D C LtA E C D" proof - have Q3: "D C OS A E" proof - have Q4: "D C TS A B" by (metis Col_perm P2 Q1 P4 bet__ts between_symmetry) then have "D C TS E B" by (metis Col_def Q1 Q2 TS_def bet__ts cong_identity invert_two_sides l9_2) thus ?thesis using OS_def Q4 by blast qed have Q5: "A C D LtA E C D" proof - have "D C A LeA D C E" proof - have "B Out D A" using P2 Q1 bet_out by auto then have "B C OS D A" by (simp add: P3 out_one_side) then have "C B OS D A" using invert_one_side by blast then have "C E OS D A" by (metis Col_def Q2 Q3 col124__nos col_one_side diff_col_ex not_col_permutation_5) then have Q5A: "A InAngle D C E" by (simp add: \C E OS D A\ Q3 invert_one_side one_side_symmetry os2__inangle) have "D C A CongA D C A" using CongA_def P6 conga_refl by auto thus ?thesis by (simp add: Q5A inangle__lea) qed then have Q6: "A C D LeA E C D" using lea_comm by blast { assume "A C D CongA E C D" then have "D C A CongA D C E" by (simp add: conga_comm) then have "C Out A E" using Q3 conga_os__out by auto then have "False" by (meson Col_def Out_cases P1 Q2 not_col_permutation_3 one_side_not_col123 out_one_side) } then have "\ A C D CongA E C D" by auto thus ?thesis by (simp add: LtA_def Q6) qed have "E C D CongA E C D" by (metis P1 P5 Q2 cong_diff conga_refl not_col_distincts) thus ?thesis using Q5 P6 conga_preserves_lta by auto qed then have "B C D LtA B D C" using Bet_cases P1 P2 Q1 Q2 bet2_lta__lta not_col_distincts by blast } then have "B C D LtA B D C" by (meson Out_def \Bet A B D \ B C D LtA B D C\ assms(1) between_symmetry) thus ?thesis by (simp add: l11_44_2_b) qed qed lemma triangle_reverse_inequality: assumes "A Out B D" and "Cong A C A D" shows "B D Le B C" proof cases assume "A Out B C" thus ?thesis by (metis assms(1) assms(2) bet__le1213 cong_pseudo_reflexivity l6_11_uniqueness l6_6 not_bet_distincts not_cong_4312) next assume "\ A Out B C" thus ?thesis using triangle_strict_reverse_inequality assms(1) assms(2) lt__le by auto qed lemma os3__lta: assumes "A B OS C D" and "B C OS A D" and "A C OS B D" shows "B A C LtA B D C" proof - have P1: "D InAngle A B C" by (simp add: assms(1) assms(2) invert_one_side os2__inangle) then obtain E where P2: "Bet A E C \ (E = B \ B Out E D)" using InAngle_def by auto have P3: "\ Col A B C" using assms(1) one_side_not_col123 by auto have P4: "\ Col A C D" using assms(3) col124__nos by auto have P5: "\ Col B C D" using assms(2) one_side_not_col124 by auto have P6: "\ Col A B D" using assms(1) one_side_not_col124 by auto { assume "E = B" then have "B A C LtA B D C" using P2 P3 bet_col by blast } { assume P7: "B Out E D" have P8: "A \ E" using P6 P7 not_col_permutation_4 out_col by blast have P9: "C \ E" using P5 P7 out_col by blast have P10: "B A C LtA B E C" proof - have P10A: "\ Col E A B" by (metis Col_def P2 P3 P8 col_transitivity_1) then have P10B: "E B A LtA B E C" using P2 P9 Tarski_neutral_dimensionless.l11_41_aux Tarski_neutral_dimensionless_axioms by fastforce have P10C: "E A B LtA B E C" using P2 P9 P10A l11_41 by auto have P11: "E A B CongA B A C" proof - have P11A: "A Out B B" using assms(2) os_distincts out_trivial by auto have "A Out C E" using P2 P8 bet_out l6_6 by auto thus ?thesis using P11A conga_right_comm out2__conga by blast qed have P12: "B E C CongA B E C" by (metis Col_def P2 P3 P9 conga_refl) thus ?thesis using P11 P10C conga_preserves_lta by auto qed have "B E C LtA B D C" proof - have U1: "E Out D B" proof - obtain pp :: "'p \ 'p \ 'p" where f1: "\p pa. p \ (pp p pa) \ pa \ (pp p pa) \ Col p pa (pp p pa)" using diff_col_ex by moura then have "\p pa pb. Col pb pa p \ \ Col pb pa (pp p pa)" by (meson l6_16_1) then have f2: "\p pa. Col pa p pa" using f1 by metis have f3: "(E = B \ D = E) \ Col D E B" using f1 by (metis Col_def P2 col_out2_col l6_16_1 out_trivial) have "\p. (A = E \ Col p A C) \ \ Col p A E" using Col_def P2 l6_16_1 by blast thus ?thesis using f3 f2 by (metis (no_types) Col_def assms(3) not_out_bet one_side_chara one_side_symmetry) qed have U2: "D \ E" using P2 P4 bet_col not_col_permutation_5 by blast have U3: "\ Col D E C" by (metis Col_def P2 P4 P9 col_transitivity_1) have U4: "Bet E D B" by (simp add: P7 U1 out2__bet) have "D C E LtA C D B" using P5 U3 U4 l11_41_aux not_col_distincts by blast have U5: "D E C LtA C D B" using P7 U4 U3 l11_41 out_diff2 by auto have "D E C CongA B E C" by (simp add: P9 U1 l6_6 out2__conga out_trivial) thus ?thesis by (metis U5 conga_preserves_lta conga_pseudo_refl lta_distincts) qed then have "B A C LtA B D C" using P10 lta_trans by blast } thus ?thesis using P2 \E = B \ B A C LtA B D C\ by blast qed lemma bet_le__lt: assumes "Bet A D B" and "A \ D" and "D \ B" and "A C Le B C" shows "D C Lt B C" proof - have P1: "A \ B" using assms(1) assms(2) between_identity by blast have "C D Lt C B" proof cases assume P3: "Col A B C" thus ?thesis proof cases assume "Bet C D B" thus ?thesis by (simp add: assms(3) bet__lt1213) next assume "\ Bet C D B" then have "D Out C B" by (metis Out_def P1 P3 assms(1) col_transitivity_2 not_col_permutation_3 not_out_bet out_col) thus ?thesis by (smt Le_cases P3 assms(1) assms(2) assms(4) bet2_le2__le bet_le_eq bet_out_1 l6_6 l6_7 nle__lt or_bet_out out2__bet out_bet__out) qed next assume Q0A: "\ Col A B C" then have Q0B: "\ Col B C D" by (meson Col_def assms(1) assms(3) col_transitivity_2) have "C B D LtA C D B" proof - have Q1: "B Out C C" using Q0A not_col_distincts out_trivial by force have Q2: "B Out A D" using Out_cases assms(1) assms(3) bet_out_1 by blast have Q3: "A Out C C" by (metis Q0A col_trivial_3 out_trivial) have Q4: "A Out B B" using P1 out_trivial by auto have "C B A LeA C A B" using Col_perm Le_cases Q0A assms(4) l11_44_2bis by blast then have T9: "C B D LeA C A B" using Q1 Q2 Q3 Q4 lea_out4__lea by blast have "C A B LtA C D B" proof - have U2: "\ Col D A C" using Q0B assms(1) assms(2) bet_col col_transitivity_2 not_col_permutation_3 not_col_permutation_4 by blast have U3: "D \ C" using Q0B col_trivial_2 by blast have U4: "D C A LtA C D B" using U2 assms(1) assms(3) l11_41_aux by auto have U5: "D A C LtA C D B" by (simp add: U2 assms(1) assms(3) l11_41) have "A Out B D" using Out_def P1 assms(1) assms(2) by auto then have "D A C CongA C A B" using Q3 conga_right_comm out2__conga by blast thus ?thesis by (metis U5 U3 assms(3) conga_preserves_lta conga_refl) qed thus ?thesis using T9 lea123456_lta__lta by blast qed thus ?thesis by (simp add: l11_44_2_b) qed thus ?thesis using Lt_cases by auto qed lemma cong2__ncol: assumes "A \ B" and "B \ C" and "A \ C" and "Cong A P B P" and "Cong A P C P" shows "\ Col A B C" proof - have "Cong B P C P" using assms(4) assms(5) cong_inner_transitivity by blast thus ?thesis using bet_le__lt by (metis assms(1) assms(2) assms(3) assms(4) assms(5) cong__le cong__nlt lt__nle not_col_permutation_5 third_point) qed lemma cong4_cop2__eq: assumes "A \ B" and "B \ C" and "A \ C" and "Cong A P B P" and "Cong A P C P" and "Coplanar A B C P" and "Cong A Q B Q" and "Cong A Q C Q" and "Coplanar A B C Q" shows "P = Q" proof - have P1: "\ Col A B C" using assms(1) assms(2) assms(3) assms(4) assms(5) cong2__ncol by auto { assume P2: "P \ Q" have P3: "\ R. Col P Q R \ (Cong A R B R \ Cong A R C R)" using P2 assms(4) assms(5) assms(7) assms(8) l4_17 not_cong_4321 by blast obtain D where P4: "D Midpoint A B" using midpoint_existence by auto have P5: "Coplanar A B C D" using P4 coplanar_perm_9 midpoint__coplanar by blast have P6: "Col P Q D" proof - have P6A: "Coplanar P Q D A" using P1 P5 assms(6) assms(9) coplanar_pseudo_trans ncop_distincts by blast then have P6B: "Coplanar P Q D B" by (metis P4 col_cop__cop midpoint_col midpoint_distinct_1) have P6D: "Cong P A P B" using assms(4) not_cong_2143 by blast have P6E: "Cong Q A Q B" using assms(7) not_cong_2143 by blast have "Cong D A D B" using Midpoint_def P4 not_cong_2134 by blast thus ?thesis using cong3_cop2__col P6A P6B assms(1) P6D P6E by blast qed obtain R1 where P7: "P \ R1 \ Q \ R1 \ D \ R1 \ Col P Q R1" using P6 diff_col_ex3 by blast obtain R2 where P8: "Bet R1 D R2 \ Cong D R2 R1 D" using segment_construction by blast have P9: "Col P Q R2" by (metis P6 P7 P8 bet_col colx) have P9A: "Cong R1 A R1 B" using P3 P7 not_cong_2143 by blast then have "Per R1 D A" using P4 Per_def by auto then have "Per A D R1" using l8_2 by blast have P10: "Cong A R1 A R2" proof - have f1: "Bet R2 D R1 \Bet R1 R2 D" by (metis (full_types) Col_def P7 P8 between_equality not_col_permutation_5) have f2: "Cong B R1 A R1" using Cong_perm \Cong R1 A R1 B\ by blast have "Cong B R1 A R2 \ Bet R1 R2 D" using f1 Cong_perm Midpoint_def P4 P8 l7_13 by blast thus ?thesis using f2 P8 bet_cong_eq cong_inner_transitivity by blast qed have "Col A B C" proof - have P11: "Cong B R1 B R2" by (metis Cong_perm P10 P3 P9 P9A cong_inner_transitivity) have P12: "Cong C R1 C R2" using P10 P3 P7 P9 cong_inner_transitivity by blast have P12A: "Coplanar A B C R1" using P2 P7 assms(6) assms(9) col_cop2__cop by blast have P12B: "Coplanar A B C R2" using P2 P9 assms(6) assms(9) col_cop2__cop by blast have "R1 \ R2" using P7 P8 between_identity by blast thus ?thesis using P10 P11 P12A P12B P12 cong3_cop2__col by blast qed then have False by (simp add: P1) } thus ?thesis by auto qed lemma t18_18_aux: assumes "Cong A B D E" and "Cong A C D F" and "F D E LtA C A B" and "\ Col A B C" and "\ Col D E F" and "D F Le D E" shows "E F Lt B C" proof - obtain G0 where P1: "C A B CongA F D G0 \ F D OS G0 E" using angle_construction_1 assms(4) assms(5) not_col_permutation_2 by blast then have P2: "\ Col F D G0" using col123__nos by auto then obtain G where P3: "D Out G0 G \ Cong D G A B" by (metis assms(4) bet_col between_trivial2 col_trivial_2 segment_construction_3) have P4: "C A B CongA F D G" proof - have P4B: "A Out C C" by (metis assms(4) col_trivial_3 out_trivial) have P4C: "A Out B B" by (metis assms(4) col_trivial_1 out_trivial) have P4D: "D Out F F" using P2 not_col_distincts out_trivial by blast have "D Out G G0" by (simp add: P3 l6_6) thus ?thesis using P1 P4B P4C P4D using l11_10 by blast qed have "D Out G G0" by (simp add: P3 l6_6) then have "D F OS G G0" using P2 not_col_permutation_4 out_one_side by blast then have "F D OS G G0" by (simp add: invert_one_side) then have P5: "F D OS G E" using P1 one_side_transitivity by blast have P6: "\ Col F D G" by (meson P5 one_side_not_col123) have P7: "Cong C B F G" using P3 P4 assms(2) cong2_conga_cong cong_commutativity cong_symmetry by blast have P8: "F E Lt F G" proof - have P9: "F D E LtA F D G" by (metis P4 assms(3) assms(5) col_trivial_1 col_trivial_3 conga_preserves_lta conga_refl) have P10: "Cong D G D E" using P3 assms(1) cong_transitivity by blast { assume P11: "Col E F G" have P12: "F D E LeA F D G" by (simp add: P9 lta__lea) have P13: "\ F D E CongA F D G" using P9 not_lta_and_conga by blast have "F D E CongA F D G" proof - have "F Out E G" using Col_cases P11 P5 col_one_side_out l6_6 by blast then have "E F D CongA G F D" by (metis assms(5) conga_os__out conga_refl l6_6 not_col_distincts one_side_reflexivity out2__conga) thus ?thesis by (meson P10 assms(2) assms(6) cong_4321 cong_inner_transitivity l11_52 le_comm) qed then have "False" using P13 by blast } then have P15: "\ Col E F G" by auto { assume P20: "Col D E G" have P21: "F D E LeA F D G" by (simp add: P9 lta__lea) have P22: "\ F D E CongA F D G" using P9 not_lta_and_conga by blast have "F D E CongA F D G" proof - have "D Out E G" by (meson Out_cases P5 P20 col_one_side_out invert_one_side not_col_permutation_5) thus ?thesis using P10 P15 \D Out G G0\ cong_inner_transitivity l6_11_uniqueness l6_7 not_col_distincts by blast qed then have "False" by (simp add: P22) } then have P16: "\ Col D E G" by auto have P17: "E InAngle F D G" using lea_in_angle by (simp add: P5 P9 lta__lea) then obtain H where P18: "Bet F H G \ (H = D \ D Out H E)" using InAngle_def by auto { assume "H = D" then have "F G E LtA F E G" using P18 P6 bet_col by blast } { assume P19: "D Out H E" have P20: "H \ F" using Out_cases P19 assms(5) out_col by blast have P21: "H \ G" using P19 P16 l6_6 out_col by blast have "F D Le G D" using P10 assms(6) cong_pseudo_reflexivity l5_6 not_cong_4312 by blast then have "H D Lt G D" using P18 P20 P21 bet_le__lt by blast then have P22: "D H Lt D E" using Lt_cases P10 cong2_lt__lt cong_reflexivity by blast then have P23: "D H Le D E \ \ Cong D H D E" using Lt_def by blast have P24: "H \ E" using P23 cong_reflexivity by blast have P25: "Bet D H E" by (simp add: P19 P23 l6_13_1) have P26: "E G OS F D" by (metis InAngle_def P15 P16 P18 P25 bet_out_1 between_symmetry in_angle_one_side not_col_distincts not_col_permutation_1) have "F G E LtA F E G" proof - have P27: "F G E LtA D E G" proof - have P28: "D G E CongA D E G" by (metis P10 P16 l11_44_1_a not_col_distincts) have "F G E LtA D G E" proof - have P29: "F G E LeA D G E" by (metis OS_def P17 P26 P5 TS_def in_angle_one_side inangle__lea_1 invert_one_side l11_24 os2__inangle) { assume "F G E CongA D G E" then have "E G F CongA E G D" by (simp add: conga_comm) then have "G Out F D" using P26 conga_os__out by auto then have "False" using P6 not_col_permutation_2 out_col by blast } then have "\ F G E CongA D G E" by auto thus ?thesis by (simp add: LtA_def P29) qed thus ?thesis by (metis P28 P6 col_trivial_3 conga_preserves_lta conga_refl) qed have "G E D LtA G E F" proof - have P30: "G E D LeA G E F" proof - have P31: "D InAngle G E F" by (simp add: P16 P17 P26 assms(5) in_angle_two_sides l11_24 not_col_permutation_5 os_ts__inangle) have "G E D CongA G E D" by (metis P16 col_trivial_1 col_trivial_2 conga_refl) thus ?thesis using P31 inangle__lea by auto qed have "\ G E D CongA G E F" by (metis OS_def P26 P5 TS_def conga_os__out invert_one_side out_col) thus ?thesis by (simp add: LtA_def P30) qed then have "D E G LtA F E G" using lta_comm by blast thus ?thesis using P27 lta_trans by blast qed } then have "F G E LtA F E G" using P18 \H = D \ F G E LtA F E G\ by blast thus ?thesis by (simp add: l11_44_2_b) qed then have "E F Lt F G" using lt_left_comm by blast thus ?thesis using P7 cong2_lt__lt cong_pseudo_reflexivity not_cong_4312 by blast qed lemma t18_18: assumes "Cong A B D E" and "Cong A C D F" and "F D E LtA C A B" shows "E F Lt B C" proof - have P1: "F \ D" using assms(3) lta_distincts by blast have P2: "E \ D" using assms(3) lta_distincts by blast have P3: "C \ A" using assms(3) lta_distincts by auto have P4: "B \ A" using assms(3) lta_distincts by blast { assume P6: "Col A B C" { assume P7: "Bet B A C" obtain C' where P8:"Bet E D C' \ Cong D C' A C" using segment_construction by blast have P9: "Cong E F E F" by (simp add: cong_reflexivity) have P10: "Cong E C' B C" using P7 P8 assms(1) l2_11_b not_cong_4321 by blast have "E F Lt E C'" proof - have P11: "Cong D F D C'" using P8 assms(2) cong_transitivity not_cong_3412 by blast have "\ Bet E D F" using Bet_perm Col_def assms(3) col_lta__out not_bet_and_out by blast thus ?thesis using P11 P8 triangle_strict_inequality by blast qed then have "E F Lt B C" using P9 P10 cong2_lt__lt by blast } { assume "\ Bet B A C" then have "E F Lt B C" using P6 assms(3) between_symmetry col_lta__bet col_permutation_2 by blast } then have "E F Lt B C" using \Bet B A C \ E F Lt B C\ by auto } { assume P12: "\ Col A B C" { assume P13: "Col D E F" { assume P14: "Bet F D E" then have "C A B LeA F D E" by (simp add: P1 P2 P3 P4 l11_31_2) then have "F D E LtA F D E" using assms(3) lea__nlta by auto then have "False" by (simp add: nlta) then have "E F Lt B C" by auto } { assume "\ Bet F D E" then have P16: "D Out F E" using P13 not_col_permutation_1 not_out_bet by blast obtain F' where P17: "A Out B F' \ Cong A F' A C" using P3 P4 segment_construction_3 by fastforce then have P18: "B F' Lt B C" by (meson P12 Tarski_neutral_dimensionless.triangle_strict_reverse_inequality Tarski_neutral_dimensionless_axioms not_cong_3412 out_col) have "Cong B F' E F" by (meson Out_cases P16 P17 assms(1) assms(2) cong_transitivity out_cong_cong) then have "E F Lt B C" using P18 cong2_lt__lt cong_reflexivity by blast } then have "E F Lt B C" using \Bet F D E \ E F Lt B C\ by blast } { assume P20: "\ Col D E F" { assume "D F Le D E" then have "E F Lt B C" by (meson P12 Tarski_neutral_dimensionless.t18_18_aux Tarski_neutral_dimensionless_axioms P20 assms(1) assms(2) assms(3)) } { assume "D E Le D F" then have "E F Lt B C" by (meson P12 P20 Tarski_neutral_dimensionless.lta_comm Tarski_neutral_dimensionless.t18_18_aux Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) lt_comm not_col_permutation_5) } then have "E F Lt B C" using \D F Le D E \ E F Lt B C\ local.le_cases by blast } then have "E F Lt B C" using \Col D E F \ E F Lt B C\ by blast } thus ?thesis using \Col A B C \ E F Lt B C\ by auto qed lemma t18_19: assumes "A \ B" and "A \ C" and "Cong A B D E" and "Cong A C D F" and "E F Lt B C" shows "F D E LtA C A B" proof - { assume P1: "C A B LeA F D E" { assume "C A B CongA F D E" then have "False" using Cong_perm assms(3) assms(4) assms(5) cong__nlt l11_49 by blast } { assume P2: "\ C A B CongA F D E" then have "C A B LtA F E D" by (metis P1 assms(3) assms(4) assms(5) cong_symmetry lea_distincts lta__nlea not_and_lt or_lta2_conga t18_18) then have "B C Lt E F" by (metis P1 P2 assms(3) assms(4) cong_symmetry lta__nlea lta_distincts or_lta2_conga t18_18) then have "False" using assms(5) not_and_lt by auto } then have "False" using \C A B CongA F D E \ False\ by auto } then have "\ C A B LeA F D E" by auto thus ?thesis using assms(1) assms(2) assms(3) assms(4) cong_identity nlea__lta by blast qed lemma acute_trivial: assumes "A \ B" shows "Acute A B A" by (metis Tarski_neutral_dimensionless.acute_distincts Tarski_neutral_dimensionless_axioms angle_partition assms l11_43) lemma acute_not_per: assumes "Acute A B C" shows "\ Per A B C" proof - obtain A' B' C' where P1: "Per A' B' C' \ A B C LtA A' B' C'" using Acute_def assms by auto thus ?thesis using acute_distincts acute_per__lta assms nlta by fastforce qed lemma angle_bisector: assumes "A \ B" and "C \ B" shows "\ P. (P InAngle A B C \ P B A CongA P B C)" proof cases assume P1: "Col A B C" thus ?thesis proof cases assume P2: "Bet A B C" then obtain Q where P3: "\ Col A B Q" using assms(1) not_col_exists by auto then obtain P where P4: "A B Perp P B \ A B OS Q P" using P1 l10_15 os_distincts by blast then have P5: "P InAngle A B C" by (metis P2 assms(2) in_angle_line os_distincts) have "P B A CongA P B C" proof - have P9: "P \ B" using P4 os_distincts by blast have "Per P B A" by (simp add: P4 Perp_perm Tarski_neutral_dimensionless.perp_per_2 Tarski_neutral_dimensionless_axioms) thus ?thesis using P2 assms(1) assms(2) P9 l11_18_1 by auto qed thus ?thesis using P5 by auto next assume T1: "\ Bet A B C" then have T2: "B Out A C" by (simp add: P1 l6_4_2) have T3: "C InAngle A B C" by (simp add: assms(1) assms(2) inangle3123) have "C B A CongA C B C" using T2 between_trivial2 l6_6 out2__conga out2_bet_out by blast thus ?thesis using T3 by auto qed next assume T4: "\ Col A B C" obtain C0 where T5: "B Out C0 C \ Cong B C0 B A" using assms(1) assms(2) l6_11_existence by fastforce obtain P where T6: "P Midpoint A C0" using midpoint_existence by auto have T6A: "\ Col A B C0" by (metis T4 T5 col3 l6_3_1 not_col_distincts out_col) have T6B: "P \ B" using Col_def Midpoint_def T6 T6A by auto have T6D: "P \ A" using T6 T6A is_midpoint_id not_col_distincts by blast have "P InAngle A B C0" using InAngle_def T5 T6 T6B assms(1) l6_3_1 midpoint_bet out_trivial by fastforce then have T7: "P InAngle A B C" using T5 T6B in_angle_trans2 l11_24 out341__inangle by blast have T8: "(P = B) \ B Out P P" using out_trivial by auto have T9: "B Out A A" by (simp add: assms(1) out_trivial) { assume T9A: "B Out P P" have "P B A CongA P B C0 \ B P A CongA B P C0 \ P A B CongA P C0 B" proof - have T9B: "Cong B P B P" by (simp add: cong_reflexivity) have T9C: "Cong B A B C0" using Cong_perm T5 by blast have "Cong P A P C0" using Midpoint_def T6 not_cong_2134 by blast thus ?thesis using l11_51 T6B assms(1) T9B T9C T6D by presburger qed then have "P B A CongA P B C0" by auto then have "P B A CongA P B C" using l11_10 T9A T9 by (meson T5 l6_6) then have "\ P. (P InAngle A B C \ P B A CongA P B C)" using T7 by auto } thus ?thesis using T6B T8 by blast qed lemma reflectl__conga: assumes "A \ B" and "B \ P" and "P P' ReflectL A B" shows "A B P CongA A B P'" proof - obtain A' where P1: "A' Midpoint P' P \ Col A B A' \ (A B Perp P' P \ P = P')" using ReflectL_def assms(3) by auto { assume P2: "A B Perp P' P" then have P3: "P \ P'" using perp_not_eq_2 by blast then have P4: "A' \ P'" using P1 is_midpoint_id by blast have P5: "A' \ P" using P1 P3 is_midpoint_id_2 by auto have "A B P CongA A B P'" proof cases assume P6: "A' = B" then have P8: "B \ P'" using P4 by auto have P9: "Per A B P" by (smt P1 P3 P6 Perp_cases col_transitivity_2 midpoint_col midpoint_distinct_1 not_col_permutation_2 perp_col2_bis perp_per_2) have "Per A B P'" by (smt Mid_cases P1 P2 P6 P8 assms(1) col_trivial_3 midpoint_col not_col_permutation_3 perp_col4 perp_per_2) thus ?thesis using l11_16 P4 P5 P6 P9 assms(1) by auto next assume T1: "A' \ B" have T2: "B A' P CongA B A' P'" proof - have T2A: "Cong B P B P'" using assms(3) col_trivial_2 is_image_spec_col_cong l10_4_spec not_cong_4321 by blast then have T2B: "A' B P CongA A' B P'" by (metis Cong_perm Midpoint_def P1 P5 T1 Tarski_neutral_dimensionless.l11_51 Tarski_neutral_dimensionless_axioms assms(2) cong_reflexivity) have "A' P B CongA A' P' B" by (simp add: P5 T2A T2B cong_reflexivity conga_comm l11_49) thus ?thesis using P5 T2A T2B cong_reflexivity l11_49 by blast qed have T3: "Cong A' B A' B" by (simp add: cong_reflexivity) have "Cong A' P A' P'" using Midpoint_def P1 not_cong_4312 by blast then have T4: "A' B P CongA A' B P' \ A' P B CongA A' P' B" using l11_49 using assms(2) T2 T3 by blast show ?thesis proof cases assume "Bet A' B A" thus ?thesis using T4 assms(1) l11_13 by blast next assume "\ Bet A' B A" then have T5: "B Out A' A" using P1 not_col_permutation_3 or_bet_out by blast have T6: "B \ P'" using T4 conga_distinct by blast have T8: "B Out A A'" by (simp add: T5 l6_6) have T9: "B Out P P" using assms(2) out_trivial by auto have "B Out P' P'" using T6 out_trivial by auto thus ?thesis using l11_10 T4 T8 T9 by blast qed qed } { assume "P = P'" then have "A B P CongA A B P'" using assms(1) assms(2) conga_refl by auto } thus ?thesis using P1 \A B Perp P' P \ A B P CongA A B P'\ by blast qed lemma conga_cop_out_reflectl__out: assumes "\ B Out A C" and "Coplanar A B C P" and "P B A CongA P B C" and "B Out A T" and "T T' ReflectL B P" shows "B Out C T'" proof - have P1: "P B T CongA P B T'" by (metis assms(3) assms(4) assms(5) conga_distinct is_image_spec_rev out_distinct reflectl__conga) have P2: "T T' Reflect B P" by (metis P1 assms(5) conga_distinct is_image_is_image_spec) have P3: "B \ T'" using CongA_def P1 by blast have P4: "P B C CongA P B T'" proof - have P5: "P B C CongA P B A" by (simp add: assms(3) conga_sym) have "P B A CongA P B T'" proof - have P7: "B Out P P" using assms(3) conga_diff45 out_trivial by blast have P8: "B Out A T" by (simp add: assms(4)) have "B Out T' T'" using P3 out_trivial by auto thus ?thesis using P1 P7 P8 l11_10 by blast qed thus ?thesis using P5 not_conga by blast qed have "P B OS C T'" proof - have P9: "P B TS A C" using assms(1) assms(2) assms(3) conga_cop__or_out_ts coplanar_perm_20 by blast then have "T \ T'" by (metis Col_perm P2 P3 TS_def assms(4) col_transitivity_2 l10_8 out_col) then have "P B TS T T'" by (metis P2 P4 conga_diff45 invert_two_sides l10_14) then have "P B TS A T'" using assms(4) col_trivial_2 out_two_sides_two_sides by blast thus ?thesis using OS_def P9 l9_2 by blast qed thus ?thesis using P4 conga_os__out by auto qed lemma col_conga_cop_reflectl__col: assumes "\ B Out A C" and "Coplanar A B C P" and "P B A CongA P B C" and "Col B A T" and "T T' ReflectL B P" shows "Col B C T'" proof cases assume "B = T" thus ?thesis using assms(5) col_image_spec__eq not_col_distincts by blast next assume P1: "B \ T" thus ?thesis proof cases assume "B Out A T" thus ?thesis using out_col conga_cop_out_reflectl__out assms(1) assms(2) assms(3) assms(5) by blast next assume P2: "\ B Out A T" obtain A' where P3: "Bet A B A' \ Cong B A' A B" using segment_construction by blast obtain C' where P4: "Bet C B C' \ Cong B C' C B" using segment_construction by blast have P5: "B Out C' T'" proof - have P6: "\ B Out A' C'" by (metis P3 P4 assms(1) between_symmetry cong_diff_2 l6_2 out_diff1 out_diff2) have P7: "Coplanar A' B C' P" proof cases assume "Col A B C" thus ?thesis by (smt P3 P4 assms(1) assms(2) assms(3) bet_col bet_neq32__neq col2_cop__cop col_transitivity_1 colx conga_diff2 conga_diff56 l6_4_2 ncoplanar_perm_15 not_col_permutation_5) next assume P7B: "\ Col A B C" have P7C: "Coplanar A B C A'" using P3 bet_col ncop__ncols by blast have P7D: "Coplanar A B C B" using ncop_distincts by blast have "Coplanar A B C C'" using P4 bet__coplanar coplanar_perm_20 by blast thus ?thesis using P7B P7C P7D assms(2) coplanar_pseudo_trans by blast qed have P8: "P B A' CongA P B C'" by (metis CongA_def P3 P4 assms(3) cong_reverse_identity conga_left_comm l11_13 not_conga_sym) have P9: "B Out A' T" by (smt Out_def P1 P2 P3 P8 assms(3) assms(4) conga_distinct l5_2 l6_4_2 not_col_permutation_4) thus ?thesis using P6 P7 P8 P9 assms(5) conga_cop_out_reflectl__out by blast qed thus ?thesis by (metis Col_def P4 col_transitivity_1 out_col out_diff1) qed qed lemma conga2_cop2__col: assumes "\ B Out A C" and "P B A CongA P B C" and "P' B A CongA P' B C" and "Coplanar A B P P'" and "Coplanar B C P P'" shows "Col B P P'" proof - obtain C' where P1: "B Out C' C \ Cong B C' B A" by (metis assms(2) conga_distinct l6_11_existence) have P1A: "Cong P A P C' \ (P \ A \ (B P A CongA B P C' \ B A P CongA B C' P))" proof - have P2: "P B A CongA P B C'" proof - have P2A: "B Out P P" using assms(2) conga_diff45 out_trivial by auto have "B Out A A" using assms(2) conga_distinct out_trivial by auto thus ?thesis using P1 P2A assms(2) l11_10 by blast qed have P3: "Cong B P B P" by (simp add: cong_reflexivity) have "Cong B A B C'" using Cong_perm P1 by blast thus ?thesis using l11_49 P2 cong_reflexivity by blast qed have P4: "P' B A CongA P' B C'" proof - have P4A: "B Out P' P'" using assms(3) conga_diff1 out_trivial by auto have "B Out A A" using assms(2) conga_distinct out_trivial by auto thus ?thesis using P1 P4A assms(3) l11_10 by blast qed have P5: "Cong B P' B P'" by (simp add: cong_reflexivity) have P5A: "Cong B A B C'" using Cong_perm P1 by blast then have P6: "P' \ A \ (B P' A CongA B P' C' \ B A P' CongA B C' P')" using P4 P5 l11_49 by blast have P7: "Coplanar B P P' A" using assms(4) ncoplanar_perm_18 by blast have P8: "Coplanar B P P' C'" by (smt Col_cases P1 assms(5) col_cop__cop ncoplanar_perm_16 ncoplanar_perm_8 out_col out_diff2) have "A \ C'" using P1 assms(1) by auto thus ?thesis using P4 P5 P7 P8 P5A P1A cong3_cop2__col l11_49 by blast qed lemma conga2_cop2__col_1: assumes "\ Col A B C" and "P B A CongA P B C" and "P' B A CongA P' B C" and "Coplanar A B C P" and "Coplanar A B C P'" shows "Col B P P'" proof - have P1: "\ B Out A C" using Col_cases assms(1) out_col by blast have P2: "Coplanar A B P P'" by (meson assms(1) assms(4) assms(5) coplanar_perm_12 coplanar_trans_1 not_col_permutation_2) have "Coplanar B C P P'" using assms(1) assms(4) assms(5) coplanar_trans_1 by auto thus ?thesis using P1 P2 conga2_cop2__col assms(2) assms(3) conga2_cop2__col by auto qed lemma col_conga__conga: assumes "P B A CongA P B C" and "Col B P P'" and "B \ P'" shows "P' B A CongA P' B C" proof cases assume "Bet P B P'" thus ?thesis using assms(1) assms(3) l11_13 by blast next assume "\ Bet P B P'" then have P1: "B Out P P'" using Col_cases assms(2) or_bet_out by blast then have P2: "B Out P' P" by (simp add: l6_6) have P3: "B Out A A" using CongA_def assms(1) out_trivial by auto have "B Out C C" using assms(1) conga_diff56 out_trivial by blast thus ?thesis using P2 P3 assms(1) l11_10 by blast qed lemma cop_inangle__ex_col_inangle: assumes "\ B Out A C" and "P InAngle A B C" and "Coplanar A B C Q" shows "\ R. (R InAngle A B C \ P \ R \ Col P Q R)" proof - have P1: "A \ B" using assms(2) inangle_distincts by blast then have P4: "A \ C" using assms(1) out_trivial by blast have P2: "C \ B" using assms(2) inangle_distincts by auto have P3: "P \ B" using InAngle_def assms(2) by auto thus ?thesis proof cases assume "P = Q" thus ?thesis using P1 P2 P4 col_trivial_1 inangle1123 inangle3123 by blast next assume P5: "P \ Q" thus ?thesis proof cases assume P6: "Col B P Q" obtain R where P7: "Bet B P R \ Cong P R B P" using segment_construction by blast have P8: "R InAngle A B C" using Out_cases P1 P2 P3 P7 assms(2) bet_out l11_25 out_trivial by blast have "P \ R" using P3 P7 cong_reverse_identity by blast thus ?thesis by (metis P3 P6 P7 P8 bet_col col_transitivity_2) next assume T1: "\ Col B P Q" thus ?thesis proof cases assume T2: "Col A B C" have T3: "Q InAngle A B C" by (metis P1 P2 T1 T2 assms(1) in_angle_line l6_4_2 not_col_distincts) thus ?thesis using P5 col_trivial_2 by blast next assume Q1: "\ Col A B C" thus ?thesis proof cases assume Q2: "Col B C P" have Q3: "\ Col B A P" using Col_perm P3 Q1 Q2 col_transitivity_2 by blast have Q4: "Coplanar B P Q A" using P2 Q2 assms(3) col2_cop__cop col_trivial_3 ncoplanar_perm_22 ncoplanar_perm_3 by blast have Q5: "Q \ P" using P5 by auto have Q6: "Col B P P" using not_col_distincts by blast have Q7: "Col Q P P" using not_col_distincts by auto have "\ Col B P A" using Col_cases Q3 by auto then obtain Q0 where P10: "Col Q P Q0 \ B P OS A Q0" using cop_not_par_same_side Q4 Q5 Q6 Q7 T1 by blast have P13: "P \ Q0" using P10 os_distincts by auto { assume "B A OS P Q0" then have ?thesis using P10 P13 assms(2) in_angle_trans not_col_permutation_4 os2__inangle by blast } { assume V1: "\ B A OS P Q0" have "\ R. Bet P R Q0 \ Col P Q R \ Col B A R" proof cases assume V3: "Col B A Q0" have "Col P Q Q0" using Col_cases P10 by auto thus ?thesis using V3 between_trivial by auto next assume V4: "\ Col B A Q0" then have V5: "\ Col Q0 B A" using Col_perm by blast have "\ Col P B A" using Col_cases Q3 by blast then obtain R where V8: "Col R B A \ Bet P R Q0" using cop_nos__ts V1 V5 by (meson P10 TS_def ncoplanar_perm_2 os__coplanar) thus ?thesis by (metis Col_def P10 P13 col_transitivity_2) qed then obtain R where V9: "Bet P R Q0 \ Col P Q R \ Col B A R" by auto have V10: "P \ R" using Q3 V9 by blast have "R InAngle A B C" proof - have W1: "\ Col B P Q0" using P10 P13 T1 col2__eq by blast have "P Out Q0 R" using V10 V9 bet_out l6_6 by auto then have "B P OS Q0 R" using Q6 W1 out_one_side_1 by blast then have "B P OS A R" using P10 one_side_transitivity by blast then have "B Out A R" using V9 col_one_side_out by auto thus ?thesis by (simp add: P2 out321__inangle) qed then have ?thesis using V10 V9 by blast } thus ?thesis using \B A OS P Q0 \ \R. R InAngle A B C \ P \ R \ Col P Q R\ by blast next assume Z1: "\ Col B C P" then have Z6: "\ Col B P C" by (simp add: not_col_permutation_5) have Z3: "Col B P P" by (simp add: col_trivial_2) have Z4: "Col Q P P" by (simp add: col_trivial_2) have "Coplanar A B C P" using Q1 assms(2) inangle__coplanar ncoplanar_perm_18 by blast then have "Coplanar B P Q C" using Q1 assms(3) coplanar_trans_1 ncoplanar_perm_5 by blast then obtain Q0 where Z5: "Col Q P Q0 \ B P OS C Q0" using cop_not_par_same_side by (metis Z3 Z4 T1 Z6) thus ?thesis proof cases assume "B C OS P Q0" thus ?thesis proof - have "\p. p InAngle C B A \ \ p InAngle C B P" using assms(2) in_angle_trans l11_24 by blast then show ?thesis by (metis Col_perm Z5 \B C OS P Q0\ l11_24 os2__inangle os_distincts) qed next assume Z6: "\ B C OS P Q0" have Z7: "\ R. Bet P R Q0 \ Col P Q R \ Col B C R" proof cases assume "Col B C Q0" thus ?thesis using Col_def Col_perm Z5 between_trivial by blast next assume Z8: "\ Col B C Q0" have "\ R. Col R B C \ Bet P R Q0" proof - have Z10: "Coplanar B C P Q0" using Z5 ncoplanar_perm_2 os__coplanar by blast have Z11: "\ Col P B C" using Col_cases Z1 by blast have "\ Col Q0 B C" using Col_perm Z8 by blast thus ?thesis using cop_nos__ts Z6 Z10 Z11 by (simp add: TS_def) qed then obtain R where "Col R B C \ Bet P R Q0" by blast thus ?thesis by (smt Z5 bet_col col2__eq col_permutation_1 os_distincts) qed then obtain R where Z12: "Bet P R Q0 \ Col P Q R \ Col B C R" by blast have Z13: "P \ R" using Z1 Z12 by auto have Z14: "\ Col B P Q0" using Z5 one_side_not_col124 by blast have "P Out Q0 R" using Z12 Z13 bet_out l6_6 by auto then have "B P OS Q0 R" using Z14 Z3 out_one_side_1 by blast then have "B P OS C R" using Z5 one_side_transitivity by blast then have "B Out C R" using Z12 col_one_side_out by blast then have "R InAngle A B C" using P1 out341__inangle by auto thus ?thesis using Z12 Z13 by auto qed qed qed qed qed qed lemma col_inangle2__out: assumes "\ Bet A B C" and "P InAngle A B C" and "Q InAngle A B C" and "Col B P Q" shows "B Out P Q" proof cases assume "Col A B C" thus ?thesis by (meson assms(1) assms(2) assms(3) assms(4) bet_in_angle_bet bet_out__bet in_angle_out l6_6 not_col_permutation_4 or_bet_out) next assume P1: "\ Col A B C" thus ?thesis proof cases assume "Col B A P" thus ?thesis by (meson assms(1) assms(2) assms(3) assms(4) bet_in_angle_bet bet_out__bet l6_6 not_col_permutation_4 or_bet_out) next assume P2: "\ Col B A P" have "\ Col B A Q" using P2 assms(3) assms(4) col2__eq col_permutation_4 inangle_distincts by blast then have "B A OS P Q" using P1 P2 assms(2) assms(3) inangle_one_side invert_one_side not_col_permutation_4 by auto thus ?thesis using assms(4) col_one_side_out by auto qed qed lemma inangle2__lea: assumes "P InAngle A B C" and "Q InAngle A B C" shows "P B Q LeA A B C" proof - have P1: "P InAngle C B A" by (simp add: assms(1) l11_24) have P2: "Q InAngle C B A" by (simp add: assms(2) l11_24) have P3: "A \ B" using assms(1) inangle_distincts by auto have P4: "C \ B" using assms(1) inangle_distincts by blast have P5: "P \ B" using assms(1) inangle_distincts by auto have P6: "Q \ B" using assms(2) inangle_distincts by auto thus ?thesis proof cases assume P7: "Col A B C" thus ?thesis proof cases assume "Bet A B C" thus ?thesis by (simp add: P3 P4 P5 P6 l11_31_2) next assume "\ Bet A B C" then have "B Out A C" using P7 not_out_bet by blast then have "B Out P Q" using Out_cases assms(1) assms(2) in_angle_out l6_7 by blast thus ?thesis by (simp add: P3 P4 l11_31_1) qed next assume T1: "\ Col A B C" thus ?thesis proof cases assume T2: "Col B P Q" have "\ Bet A B C" using T1 bet_col by auto then have "B Out P Q" using T2 assms(1) assms(2) col_inangle2__out by auto thus ?thesis by (simp add: P3 P4 l11_31_1) next assume T3: "\ Col B P Q" thus ?thesis proof cases assume "Col B A P" then have "B Out A P" using Col_def T1 assms(1) col_in_angle_out by blast then have "P B Q CongA A B Q" using P6 out2__conga out_trivial by auto thus ?thesis using LeA_def assms(2) by blast next assume W0: "\ Col B A P" show ?thesis proof cases assume "Col B C P" then have "B Out C P" by (metis P1 P3 T1 bet_out_1 col_in_angle_out out_col) thus ?thesis by (smt P3 P4 P6 Tarski_neutral_dimensionless.lea_left_comm Tarski_neutral_dimensionless.lea_out4__lea Tarski_neutral_dimensionless_axioms assms(2) inangle__lea_1 out_trivial) next assume W0A: "\ Col B C P" show ?thesis proof cases assume "Col B A Q" then have "B Out A Q" using Col_def T1 assms(2) col_in_angle_out by blast thus ?thesis by (smt P3 P4 P5 Tarski_neutral_dimensionless.lea_left_comm Tarski_neutral_dimensionless.lea_out4__lea Tarski_neutral_dimensionless_axioms assms(1) inangle__lea out_trivial) next assume W0AA: "\ Col B A Q" thus ?thesis proof cases assume "Col B C Q" then have "B Out C Q" using Bet_cases P2 T1 bet_col col_in_angle_out by blast thus ?thesis by (smt P1 P3 P4 P5 Tarski_neutral_dimensionless.lea_comm Tarski_neutral_dimensionless.lea_out4__lea Tarski_neutral_dimensionless_axioms inangle__lea out_trivial) next assume W0B: "\ Col B C Q" have W1: "Coplanar B P A Q" by (metis Col_perm T1 assms(1) assms(2) col__coplanar inangle_one_side ncoplanar_perm_13 os__coplanar) have W2: "\ Col A B P" by (simp add: W0 not_col_permutation_4) have W3: "\ Col Q B P" using Col_perm T3 by blast then have W4: "B P TS A Q \ B P OS A Q" using cop__one_or_two_sides by (simp add: W1 W2) { assume W4A: "B P TS A Q" have "Q InAngle P B C" proof - have W5: "P B OS C Q" using OS_def P1 W0 W0A W4A in_angle_two_sides invert_two_sides l9_2 by blast have "C B OS P Q" by (meson P1 P2 T1 W0A W0B inangle_one_side not_col_permutation_3 not_col_permutation_4) thus ?thesis by (simp add: W5 invert_one_side os2__inangle) qed then have "P B Q LeA A B C" by (meson assms(1) inangle__lea inangle__lea_1 lea_trans) } { assume W6: "B P OS A Q" have "B A OS P Q" using Col_perm T1 W2 W0AA assms(1) assms(2) inangle_one_side invert_one_side by blast then have "Q InAngle P B A" by (simp add: W6 os2__inangle) then have "P B Q LeA A B C" by (meson P1 inangle__lea inangle__lea_1 lea_right_comm lea_trans) } thus ?thesis using W4 \B P TS A Q \ P B Q LeA A B C\ by blast qed qed qed qed qed qed qed lemma conga_inangle_per__acute: assumes "Per A B C" and "P InAngle A B C" and "P B A CongA P B C" shows "Acute A B P" proof - have P1: "\ Col A B C" using assms(1) assms(3) conga_diff2 conga_diff56 l8_9 by blast have P2: "A B P LeA A B C" by (simp add: assms(2) inangle__lea) { assume "A B P CongA A B C" then have P3: "Per A B P" by (meson Tarski_neutral_dimensionless.l11_17 Tarski_neutral_dimensionless.not_conga_sym Tarski_neutral_dimensionless_axioms assms(1)) have P4: "Coplanar P C A B" using assms(2) inangle__coplanar ncoplanar_perm_3 by blast have P5: "P \ B" using assms(2) inangle_distincts by blast have "Per C B P" using P3 Per_cases assms(3) l11_17 by blast then have "False" using P1 P3 P4 P5 col_permutation_1 cop_per2__col by blast } then have "\ A B P CongA A B C" by auto then have "A B P LtA A B C" by (simp add: LtA_def P2) thus ?thesis using Acute_def assms(1) by blast qed lemma conga_inangle2_per__acute: assumes "Per A B C" and "P InAngle A B C" and "P B A CongA P B C" and "Q InAngle A B C" shows "Acute P B Q" proof - have P1: "P InAngle C B A" using assms(2) l11_24 by auto have P2: "Q InAngle C B A" using assms(4) l11_24 by blast have P3: "A \ B" using assms(3) conga_diff2 by auto have P5: "P \ B" using assms(2) inangle_distincts by blast have P7: "\ Col A B C" using assms(1) assms(3) conga_distinct l8_9 by blast have P8: "Acute A B P" using assms(1) assms(2) assms(3) conga_inangle_per__acute by auto { assume "Col P B A" then have "Col P B C" using assms(3) col_conga_col by blast then have "False" using Col_perm P5 P7 \Col P B A\ col_transitivity_2 by blast } then have P9: "\ Col P B A" by auto have P10: "\ Col P B C" using \Col P B A \ False\ assms(3) ncol_conga_ncol by blast have P11: "\ Bet A B C" using P7 bet_col by blast show ?thesis proof cases assume "Col B A Q" then have "B Out A Q" using P11 assms(4) col_in_angle_out by auto thus ?thesis using Out_cases P5 P8 acute_out2__acute acute_sym out_trivial by blast next assume S0: "\ Col B A Q" show ?thesis proof cases assume S1: "Col B C Q" then have "B Out C Q" using P11 P2 between_symmetry col_in_angle_out by blast then have S2: "B Out Q C" using l6_6 by blast have S3: "B Out P P" by (simp add: P5 out_trivial) have "B Out A A" by (simp add: P3 out_trivial) then have "A B P CongA P B Q" using S2 conga_left_comm l11_10 S3 assms(3) by blast thus ?thesis using P8 acute_conga__acute by blast next assume S4: "\ Col B C Q" show ?thesis proof cases assume "Col B P Q" thus ?thesis using out__acute col_inangle2__out P11 assms(2) assms(4) by blast next assume S5: "\ Col B P Q" have S6: "Coplanar B P A Q" by (metis Col_perm P7 assms(2) assms(4) coplanar_trans_1 inangle__coplanar ncoplanar_perm_12 ncoplanar_perm_21) have S7: "\ Col A B P" using Col_cases P9 by auto have "\ Col Q B P" using Col_perm S5 by blast then have S8: "B P TS A Q \ B P OS A Q" using cop__one_or_two_sides S6 S7 by blast { assume S9: "B P TS A Q" have S10: "Acute P B C" using P8 acute_conga__acute acute_sym assms(3) by blast have "Q InAngle P B C" proof - have S11: "P B OS C Q" by (metis Col_perm OS_def P1 P10 P9 S9 in_angle_two_sides invert_two_sides l9_2) have "C B OS P Q" by (meson P1 P10 P2 P7 S4 inangle_one_side not_col_permutation_3 not_col_permutation_4) thus ?thesis by (simp add: S11 invert_one_side os2__inangle) qed then have "P B Q LeA P B C" by (simp add: inangle__lea) then have "Acute P B Q" using S10 acute_lea_acute by blast } { assume S12: "B P OS A Q" have "B A OS P Q" using Col_perm P7 S7 S0 assms(2) assms(4) inangle_one_side invert_one_side by blast then have "Q InAngle P B A" by (simp add: S12 os2__inangle) then have "Q B P LeA P B A" by (simp add: P3 P5 inangle1123 inangle2__lea) then have "P B Q LeA A B P" by (simp add: lea_comm) then have "Acute P B Q" using P8 acute_lea_acute by blast } thus ?thesis using \B P TS A Q \ Acute P B Q\ S8 by blast qed qed qed qed lemma lta_os__ts: assumes (*"\ Col A O1 P" and*) "A O1 P LtA A O1 B" and "O1 A OS B P" shows "O1 P TS A B" proof - have "A O1 P LeA A O1 B" by (simp add: assms(1) lta__lea) then have "\ P0. P0 InAngle A O1 B \ A O1 P CongA A O1 P0" by (simp add: LeA_def) then obtain P' where P1: "P' InAngle A O1 B \ A O1 P CongA A O1 P'" by blast have P2: "\ Col A O1 B" using assms(2) col123__nos not_col_permutation_4 by blast obtain R where P3: "O1 A TS B R \ O1 A TS P R" using OS_def assms(2) by blast { assume "Col B O1 P" then have "Bet B O1 P" by (metis Tarski_neutral_dimensionless.out2__conga Tarski_neutral_dimensionless_axioms assms(1) assms(2) between_trivial col_trivial_2 lta_not_conga one_side_chara or_bet_out out_trivial) then have "O1 A TS B P" using assms(2) col_trivial_1 one_side_chara by blast then have P6: "\ O1 A OS B P" using l9_9_bis by auto then have "False" using P6 assms(2) by auto } then have P4: "\ Col B O1 P" by auto thus ?thesis by (meson P3 assms(1) inangle__lta l9_8_1 not_and_lta not_col_permutation_4 os_ts__inangle two_sides_cases) qed lemma bet__suppa: assumes "A \ B" and "B \ C" and "B \ A'" and "Bet A B A'" shows "A B C SuppA C B A'" proof - have "C B A' CongA C B A'" using assms(2) assms(3) conga_refl by auto thus ?thesis using assms(4) assms(1) SuppA_def by auto qed lemma ex_suppa: assumes "A \ B" and "B \ C" shows "\ D E F. A B C SuppA D E F" proof - obtain A' where "Bet A B A' \ Cong B A' A B" using segment_construction by blast thus ?thesis by (meson assms(1) assms(2) bet__suppa point_construction_different) qed lemma suppa_distincts: assumes "A B C SuppA D E F" shows "A \ B \ B \ C \ D \ E \ E \ F" using CongA_def SuppA_def assms by auto lemma suppa_right_comm: assumes "A B C SuppA D E F" shows "A B C SuppA F E D" using SuppA_def assms conga_left_comm by auto lemma suppa_left_comm: assumes "A B C SuppA D E F" shows "C B A SuppA D E F" proof - obtain A' where P1: "Bet A B A' \ D E F CongA C B A'" using SuppA_def assms by auto obtain C' where P2: "Bet C B C' \ Cong B C' C B" using segment_construction by blast then have "C B A' CongA A B C'" by (metis Bet_cases P1 SuppA_def assms cong_diff_3 conga_diff45 conga_diff56 conga_left_comm l11_14) then have "D E F CongA A B C'" using P1 conga_trans by blast thus ?thesis by (metis CongA_def P1 P2 SuppA_def) qed lemma suppa_comm: assumes "A B C SuppA D E F" shows "C B A SuppA F E D" using assms suppa_left_comm suppa_right_comm by blast lemma suppa_sym: assumes "A B C SuppA D E F" shows "D E F SuppA A B C" proof - obtain A' where P1: "Bet A B A' \ D E F CongA C B A'" using SuppA_def assms by auto obtain D' where P2: "Bet D E D' \ Cong E D' D E" using segment_construction by blast have "A' B C CongA D E F" using P1 conga_right_comm not_conga_sym by blast then have "A B C CongA F E D'" by (metis P1 P2 Tarski_neutral_dimensionless.conga_right_comm Tarski_neutral_dimensionless.l11_13 Tarski_neutral_dimensionless.suppa_distincts Tarski_neutral_dimensionless_axioms assms between_symmetry cong_diff_3) thus ?thesis by (metis CongA_def P1 P2 SuppA_def) qed lemma conga2_suppa__suppa: assumes "A B C CongA A' B' C'" and "D E F CongA D' E' F'" and "A B C SuppA D E F" shows "A' B' C' SuppA D' E' F'" proof - obtain A0 where P1: "Bet A B A0 \ D E F CongA C B A0" using SuppA_def assms(3) by auto then have "A B C SuppA D' E' F'" by (metis Tarski_neutral_dimensionless.SuppA_def Tarski_neutral_dimensionless_axioms assms(2) assms(3) conga_sym conga_trans) then have P2: "D' E' F' SuppA A B C" by (simp add: suppa_sym) then obtain D0 where P3: "Bet D' E' D0 \ A B C CongA F' E' D0" using P2 SuppA_def by auto have P5: "A' B' C' CongA F' E' D0" using P3 assms(1) not_conga not_conga_sym by blast then have "D' E' F' SuppA A' B' C'" using P2 P3 SuppA_def by auto thus ?thesis by (simp add: suppa_sym) qed lemma suppa2__conga456: assumes "A B C SuppA D E F" and "A B C SuppA D' E' F'" shows "D E F CongA D' E' F'" proof - obtain A' where P1: "Bet A B A' \ D E F CongA C B A'" using SuppA_def assms(1) by auto obtain A'' where P2: "Bet A B A'' \ D' E' F' CongA C B A''" using SuppA_def assms(2) by auto have "C B A' CongA C B A''" proof - have P3: "B Out C C" using P1 by (simp add: CongA_def out_trivial) have "B Out A'' A'" using P1 P2 l6_2 by (metis assms(1) between_symmetry conga_distinct suppa_distincts) thus ?thesis by (simp add: P3 out2__conga) qed then have "C B A' CongA D' E' F'" using P2 not_conga not_conga_sym by blast thus ?thesis using P1 not_conga by blast qed lemma suppa2__conga123: assumes "A B C SuppA D E F" and "A' B' C' SuppA D E F" shows "A B C CongA A' B' C'" using assms(1) assms(2) suppa2__conga456 suppa_sym by blast lemma bet_out__suppa: assumes "A \ B" and "B \ C" and "Bet A B C" and "E Out D F" shows "A B C SuppA D E F" proof - have "D E F CongA C B C" using assms(2) assms(4) l11_21_b out_trivial by auto thus ?thesis using SuppA_def assms(1) assms(3) by blast qed lemma bet_suppa__out: assumes "Bet A B C" and "A B C SuppA D E F" shows "E Out D F" proof - have "A B C SuppA C B C" using assms(1) assms(2) bet__suppa suppa_distincts by auto then have "C B C CongA D E F" using assms(2) suppa2__conga456 by auto thus ?thesis using eq_conga_out by auto qed lemma out_suppa__bet: assumes "B Out A C" and "A B C SuppA D E F" shows "Bet D E F" proof - obtain B' where P1: "Bet A B B' \ Cong B B' A B" using segment_construction by blast have "A B C SuppA A B B'" by (metis P1 assms(1) assms(2) bet__suppa bet_cong_eq bet_out__bet suppa_distincts suppa_left_comm) then have "A B B' CongA D E F" using assms(2) suppa2__conga456 by auto thus ?thesis using P1 bet_conga__bet by blast qed lemma per_suppa__per: assumes "Per A B C" and "A B C SuppA D E F" shows "Per D E F" proof - obtain A' where P1: "Bet A B A' \ D E F CongA C B A'" using SuppA_def assms(2) by auto have "Per C B A'" proof - have P2: "A \ B" using assms(2) suppa_distincts by auto have P3: "Per C B A" by (simp add: assms(1) l8_2) have "Col B A A'" using P1 Col_cases Col_def by blast thus ?thesis by (metis P2 P3 per_col) qed thus ?thesis using P1 l11_17 not_conga_sym by blast qed lemma per2__suppa: assumes "A \ B" and "B \ C" and "D \ E" and "E \ F" and "Per A B C" and "Per D E F" shows "A B C SuppA D E F" proof - obtain D' E' F' where P1: "A B C SuppA D' E' F'" using assms(1) assms(2) ex_suppa by blast have "D' E' F' CongA D E F" using P1 assms(3) assms(4) assms(5) assms(6) l11_16 per_suppa__per suppa_distincts by blast thus ?thesis by (meson P1 conga2_suppa__suppa suppa2__conga123) qed lemma suppa__per: assumes "A B C SuppA A B C" shows "Per A B C" proof - obtain A' where P1: "Bet A B A' \ A B C CongA C B A'" using SuppA_def assms by auto then have "C B A CongA C B A'" by (simp add: conga_left_comm) thus ?thesis using P1 Per_perm l11_18_2 by blast qed lemma acute_suppa__obtuse: assumes "Acute A B C" and "A B C SuppA D E F" shows "Obtuse D E F" proof - obtain A' where P1: "Bet A B A' \ D E F CongA C B A'" using SuppA_def assms(2) by auto then have "Obtuse C B A'" by (metis Tarski_neutral_dimensionless.obtuse_sym Tarski_neutral_dimensionless_axioms acute_bet__obtuse assms(1) conga_distinct) thus ?thesis by (meson P1 Tarski_neutral_dimensionless.conga_obtuse__obtuse Tarski_neutral_dimensionless.not_conga_sym Tarski_neutral_dimensionless_axioms) qed lemma obtuse_suppa__acute: assumes "Obtuse A B C" and "A B C SuppA D E F" shows "Acute D E F" proof - obtain A' where P1: "Bet A B A' \ D E F CongA C B A'" using SuppA_def assms(2) by auto then have "Acute C B A'" using acute_sym assms(1) bet_obtuse__acute conga_distinct by blast thus ?thesis using P1 acute_conga__acute not_conga_sym by blast qed lemma lea_suppa2__lea: assumes "A B C SuppA A' B' C'" and "D E F SuppA D' E' F'" "A B C LeA D E F" shows "D' E' F' LeA A' B' C'" proof - obtain A0 where P1: "Bet A B A0 \ A' B' C' CongA C B A0" using SuppA_def assms(1) by auto obtain D0 where P2: "Bet D E D0 \ D' E' F' CongA F E D0" using SuppA_def assms(2) by auto have "F E D0 LeA C B A0" proof - have P3: "D0 \ E" using CongA_def P2 by auto have P4: "A0 \ B" using CongA_def P1 by blast have P6: "Bet D0 E D" by (simp add: P2 between_symmetry) have "Bet A0 B A" by (simp add: P1 between_symmetry) thus ?thesis by (metis P3 P4 P6 assms(3) l11_36_aux2 lea_comm lea_distincts) qed thus ?thesis by (meson P1 P2 Tarski_neutral_dimensionless.l11_30 Tarski_neutral_dimensionless.not_conga_sym Tarski_neutral_dimensionless_axioms) qed lemma lta_suppa2__lta: assumes "A B C SuppA A' B' C'" and "D E F SuppA D' E' F'" and "A B C LtA D E F" shows "D' E' F' LtA A' B' C'" proof - obtain A0 where P1: "Bet A B A0 \ A' B' C' CongA C B A0" using SuppA_def assms(1) by auto obtain D0 where P2: "Bet D E D0 \ D' E' F' CongA F E D0" using SuppA_def assms(2) by auto have "F E D0 LtA C B A0" proof - have P5: "A0 \ B" using CongA_def P1 by blast have "D0 \ E" using CongA_def P2 by auto thus ?thesis using assms(3) P1 P5 P2 bet2_lta__lta lta_comm by blast qed thus ?thesis using P1 P2 conga_preserves_lta not_conga_sym by blast qed lemma suppa_dec: "A B C SuppA D E F \ \ A B C SuppA D E F" by simp lemma acute_one_side_aux: assumes "C A OS P B" and "Acute A C P" and "C A Perp B C" shows "C B OS A P" proof - obtain R where T1: "C A TS P R \ C A TS B R" using OS_def assms(1) by blast obtain A' B' C' where P1: "Per A' B' C' \ A C P LtA A' B' C'" using Acute_def assms(2) by auto have P2: "Per A C B" by (simp add: assms(3) perp_per_1) then have P3: "A' B' C' CongA A C B" using P1 assms(1) l11_16 lta_distincts os_distincts by blast have P4: "A C P LtA A C B" by (metis P2 acute_per__lta assms(1) assms(2) os_distincts) { assume P4A: "Col P C B" have "Per A C P" proof - have P4B: "C \ B" using assms(1) os_distincts by blast have P4C: "Per A C B" by (simp add: P2) have "Col C B P" using P4A Col_cases by auto thus ?thesis using per_col P4B P4C by blast qed then have "False" using acute_not_per assms(2) by auto } then have P5: "\ Col P C B" by auto have P6: "\ Col A C P" using assms(1) col123__nos not_col_permutation_4 by blast have P7: "C B TS A P \ C B OS A P" using P5 assms(1) not_col_permutation_4 os_ts1324__os two_sides_cases by blast { assume P8: "C B TS A P" then obtain T where P9: "Col T C B \ Bet A T P" using TS_def by blast then have P10: "C \ T" using Col_def P6 P9 by auto have "T InAngle A C P" by (meson P4 P5 P8 Tarski_neutral_dimensionless.inangle__lta Tarski_neutral_dimensionless_axioms assms(1) not_and_lta not_col_permutation_3 os_ts__inangle) then have "C A OS T P" by (metis P10 P9 T1 TS_def col123__nos in_angle_one_side invert_one_side l6_16_1 one_side_reflexivity) then have P13: "C A OS T B" using assms(1) one_side_transitivity by blast have "C B OS A P" by (meson P4 Tarski_neutral_dimensionless.lta_os__ts Tarski_neutral_dimensionless_axioms assms(1) one_side_symmetry os_ts1324__os) } thus ?thesis using P7 by blast qed lemma acute_one_side_aux0: assumes "Col A C P" and "Acute A C P" and "C A Perp B C" shows "C B OS A P" proof - have "Per A C B" by (simp add: assms(3) perp_per_1) then have P1: "A C P LtA A C B" using Tarski_neutral_dimensionless.acute_per__lta Tarski_neutral_dimensionless_axioms acute_distincts assms(2) assms(3) perp_not_eq_2 by fastforce have P2: "C Out A P" using acute_col__out assms(1) assms(2) by auto thus ?thesis using Perp_cases assms(3) out_one_side perp_not_col by blast qed lemma acute_cop_perp__one_side: assumes "Acute A C P" and "C A Perp B C" and "Coplanar A B C P" shows "C B OS A P" proof cases assume "Col A C P" thus ?thesis by (simp add: acute_one_side_aux0 assms(1) assms(2)) next assume P1: "\ Col A C P" have P2: "C A TS P B \ C A OS P B" using Col_cases P1 assms(2) assms(3) cop_nos__ts coplanar_perm_13 perp_not_col by blast { assume P3: "C A TS P B" obtain Bs where P4: "C Midpoint B Bs" using symmetric_point_construction by auto have "C A TS Bs B" by (metis P3 P4 assms(2) bet__ts l9_2 midpoint_bet midpoint_distinct_2 perp_not_col ts_distincts) then have P6: "C A OS P Bs" using P3 l9_8_1 by auto have "C Bs Perp A C" proof - have P6A: "C \ Bs" using P6 os_distincts by blast have "Col C B Bs" using Bet_cases Col_def P4 midpoint_bet by blast thus ?thesis using Perp_cases P6A assms(2) perp_col by blast qed then have "Bs C Perp C A" using Perp_perm by blast then have "C A Perp Bs C" using Perp_perm by blast then have "C B OS A P" using acute_one_side_aux by (metis P4 P6 assms(1) assms(2) col_one_side midpoint_col not_col_permutation_5 perp_distinct) } { assume "C A OS P B" then have "C B OS A P" using acute_one_side_aux using assms(1) assms(2) by blast } thus ?thesis using P2 \C A TS P B \ C B OS A P\ by auto qed lemma acute__not_obtuse: assumes "Acute A B C" shows "\ Obtuse A B C" using acute_obtuse__lta assms nlta by blast subsubsection "Sum of angles" lemma suma_distincts: assumes "A B C D E F SumA G H I" shows "A \ B \ B \C \ D \ E \ E \ F \ G \ H \ H \ I" proof - obtain J where "C B J CongA D E F \ \ B C OS A J \ Coplanar A B C J \ A B J CongA G H I" using SumA_def assms by auto thus ?thesis using CongA_def by blast qed lemma trisuma_distincts: assumes "A B C TriSumA D E F" shows "A \ B \ B \ C \ A \ C \ D \ E \ E \ F" proof - obtain G H I where "A B C B C A SumA G H I \ G H I C A B SumA D E F" using TriSumA_def assms by auto thus ?thesis using suma_distincts by blast qed lemma ex_suma: assumes "A \ B" and "B \ C" and "D \ E" and "E \ F" shows "\ G H I. A B C D E F SumA G H I" proof - have "\ I. A B C D E F SumA A B I" proof cases assume P1: "Col A B C" obtain J where P2: "D E F CongA C B J \ Coplanar C B J A" using angle_construction_4 using assms(2) assms(3) assms(4) by presburger have P3: "J \ B" using CongA_def P2 by blast have "\ B C OS A J" by (metis P1 between_trivial2 one_side_chara) then have "A B C D E F SumA A B J" by (meson P2 P3 SumA_def assms(1) conga_refl ncoplanar_perm_15 not_conga_sym) thus ?thesis by blast next assume T1: "\ Col A B C" show ?thesis proof cases assume T2: "Col D E F" show ?thesis proof cases assume T3: "Bet D E F" obtain J where T4: "B Midpoint C J" using symmetric_point_construction by blast have "A B C D E F SumA A B J" proof - have "C B J CongA D E F" by (metis T3 T4 assms(2) assms(3) assms(4) conga_line midpoint_bet midpoint_distinct_2) moreover have "\ B C OS A J" by (simp add: T4 col124__nos midpoint_col) moreover have "Coplanar A B C J" using T3 bet__coplanar bet_conga__bet calculation(1) conga_sym ncoplanar_perm_15 by blast moreover have "A B J CongA A B J" using CongA_def assms(1) calculation(1) conga_refl by auto ultimately show ?thesis using SumA_def by blast qed then show ?thesis by auto next assume T5: "\ Bet D E F" have "A B C D E F SumA A B C" proof - have "E Out D F" using T2 T5 l6_4_2 by auto then have "C B C CongA D E F" using assms(2) l11_21_b out_trivial by auto moreover have "\ B C OS A C" using os_distincts by blast moreover have "Coplanar A B C C" using ncop_distincts by auto moreover have "A B C CongA A B C" using assms(1) assms(2) conga_refl by auto ultimately show ?thesis using SumA_def by blast qed then show ?thesis by auto qed next assume T6: "\ Col D E F" then obtain J where T7: "D E F CongA C B J \ C B TS J A" using T1 ex_conga_ts not_col_permutation_4 not_col_permutation_5 by presburger then show ?thesis proof - have "C B J CongA D E F" using T7 not_conga_sym by blast moreover have "\ B C OS A J" by (simp add: T7 invert_two_sides l9_2 l9_9) moreover have "Coplanar A B C J" using T7 ncoplanar_perm_15 ts__coplanar by blast moreover have "A B J CongA A B J" using T7 assms(1) conga_diff56 conga_refl by blast ultimately show ?thesis using SumA_def by blast qed qed qed then show ?thesis by auto qed lemma suma2__conga: assumes "A B C D E F SumA G H I" and "A B C D E F SumA G' H' I'" shows "G H I CongA G' H' I'" proof - obtain J where P1: "C B J CongA D E F \ \ B C OS A J \ Coplanar A B C J \ A B J CongA G H I" using SumA_def assms(1) by blast obtain J' where P2: "C B J' CongA D E F \ \ B C OS A J' \ Coplanar A B C J' \ A B J' CongA G' H' I'" using SumA_def assms(2) by blast have P3: "C B J CongA C B J'" proof - have "C B J CongA D E F" by (simp add: P1) moreover have "D E F CongA C B J'" by (simp add: P2 conga_sym) ultimately show ?thesis using not_conga by blast qed have P4: "A B J CongA A B J'" proof cases assume P5: "Col A B C" then show ?thesis proof cases assume P6: "Bet A B C" show ?thesis proof - have "C B J CongA C B J'" by (simp add: P3) moreover have "Bet C B A" by (simp add: P6 between_symmetry) moreover have "A \ B" using assms(1) suma_distincts by blast ultimately show ?thesis using l11_13 by blast qed next assume P7: "\ Bet A B C" moreover have "B Out A C" by (simp add: P5 calculation l6_4_2) moreover have "B \ J" using CongA_def P3 by blast then moreover have "B Out J J" using out_trivial by auto moreover have "B \ J'" using CongA_def P3 by blast then moreover have "B Out J' J'" using out_trivial by auto ultimately show ?thesis using P3 l11_10 by blast qed next assume P8: "\ Col A B C" show ?thesis proof cases assume P9: "Col D E F" have "B Out J' J" proof cases assume P10: "Bet D E F" show ?thesis proof - have "D E F CongA J' B C" using P2 conga_right_comm not_conga_sym by blast then have "Bet J' B C" using P10 bet_conga__bet by blast moreover have "D E F CongA J B C" by (simp add: P1 conga_right_comm conga_sym) then moreover have "Bet J B C" using P10 bet_conga__bet by blast ultimately show ?thesis by (metis CongA_def P3 l6_2) qed next assume P11: "\ Bet D E F" have P12: "E Out D F" by (simp add: P11 P9 l6_4_2) show ?thesis proof - have "B Out J' C" proof - have "D E F CongA J' B C" using P2 conga_right_comm conga_sym by blast then show ?thesis using l11_21_a P12 by blast qed moreover have "B Out C J" by (metis P3 P8 bet_conga__bet calculation col_conga_col col_out2_col l6_4_2 l6_6 not_col_distincts not_conga_sym out_bet_out_1 out_trivial) ultimately show ?thesis using l6_7 by blast qed qed then show ?thesis using P8 not_col_distincts out2__conga out_trivial by blast next assume P13: "\ Col D E F" show ?thesis proof - have "B C TS A J" proof - have "Coplanar B C A J" using P1 coplanar_perm_8 by blast moreover have "\ Col A B C" by (simp add: P8) moreover have "\ B C OS A J" using P1 by simp moreover have "\ Col J B C" proof - have "D E F CongA J B C" using P1 conga_right_comm not_conga_sym by blast then show ?thesis using P13 ncol_conga_ncol by blast qed ultimately show ?thesis using cop__one_or_two_sides by blast qed moreover have "B C TS A J'" proof - have "Coplanar B C A J'" using P2 coplanar_perm_8 by blast moreover have "\ Col A B C" by (simp add: P8) moreover have "\ B C OS A J'" using P2 by simp moreover have "\ Col J' B C" proof - have "D E F CongA J' B C" using P2 conga_right_comm not_conga_sym by blast then show ?thesis using P13 ncol_conga_ncol by blast qed ultimately show ?thesis using cop_nos__ts by blast qed moreover have "A B C CongA A B C" by (metis P8 conga_pseudo_refl conga_right_comm not_col_distincts) moreover have "C B J CongA C B J'" by (simp add: P3) ultimately show ?thesis using l11_22a by blast qed qed qed then show ?thesis by (meson P1 P2 not_conga not_conga_sym) qed lemma suma_sym: assumes "A B C D E F SumA G H I" shows "D E F A B C SumA G H I" proof - obtain J where P1: "C B J CongA D E F \ \ B C OS A J \ Coplanar A B C J \ A B J CongA G H I" using SumA_def assms(1) by blast show ?thesis proof cases assume P2: "Col A B C" then show ?thesis proof cases assume P3: "Bet A B C" obtain K where P4: "Bet F E K \ Cong F E E K" using Cong_perm segment_construction by blast show ?thesis proof - have P5: "F E K CongA A B C" by (metis CongA_def P1 P3 P4 cong_diff conga_line) moreover have "\ E F OS D K" using P4 bet_col col124__nos invert_one_side by blast moreover have "Coplanar D E F K" using P4 bet__coplanar ncoplanar_perm_15 by blast moreover have "D E K CongA G H I" proof - have "D E K CongA A B J" proof - have "F E D CongA C B J" by (simp add: P1 conga_left_comm conga_sym) moreover have "Bet F E K" by (simp add: P4) moreover have "K \ E" using P4 calculation(1) cong_identity conga_diff1 by blast moreover have "Bet C B A" by (simp add: Bet_perm P3) moreover have "A \ B" using CongA_def P5 by blast ultimately show ?thesis using conga_right_comm l11_13 not_conga_sym by blast qed then show ?thesis using P1 not_conga by blast qed ultimately show ?thesis using SumA_def by blast qed next assume T1: "\ Bet A B C" then have T2: "B Out A C" by (simp add: P2 l6_4_2) show ?thesis proof - have "F E F CongA A B C" by (metis T2 assms l11_21_b out_trivial suma_distincts) moreover have "\ E F OS D F" using os_distincts by auto moreover have "Coplanar D E F F" using ncop_distincts by auto moreover have "D E F CongA G H I" proof - have "A B J CongA D E F" proof - have "C B J CongA D E F" by (simp add: P1) moreover have "B Out A C" by (simp add: T2) moreover have "J \ B" using calculation(1) conga_distinct by auto moreover have "D \ E" using calculation(1) conga_distinct by blast moreover have "F \ E" using calculation(1) conga_distinct by blast ultimately show ?thesis by (meson Out_cases not_conga out2__conga out_trivial) qed then have "D E F CongA A B J" using not_conga_sym by blast then show ?thesis using P1 not_conga by blast qed ultimately show ?thesis using SumA_def by blast qed qed next assume Q1: "\ Col A B C" show ?thesis proof cases assume Q2: "Col D E F" obtain K where Q3: "A B C CongA F E K" using P1 angle_construction_3 conga_diff1 conga_diff56 by fastforce show ?thesis proof - have "F E K CongA A B C" by (simp add: Q3 conga_sym) moreover have "\ E F OS D K" using Col_cases Q2 one_side_not_col123 by blast moreover have "Coplanar D E F K" by (simp add: Q2 col__coplanar) moreover have "D E K CongA G H I" proof - have "D E K CongA A B J" proof cases assume "Bet D E F" then have "J B A CongA D E K" by (metis P1 bet_conga__bet calculation(1) conga_diff45 conga_right_comm l11_13 not_conga_sym) then show ?thesis using conga_right_comm not_conga_sym by blast next assume "\ Bet D E F" then have W2: "E Out D F" using Q2 or_bet_out by blast have "A B J CongA D E K" proof - have "A B C CongA F E K" by (simp add: Q3) moreover have "A \ B" using Q1 col_trivial_1 by auto moreover have "E Out D F" by (simp add: W2) moreover have "B Out J C" proof - have "D E F CongA J B C" by (simp add: P1 conga_left_comm conga_sym) then show ?thesis using W2 out_conga_out by blast qed moreover have "K \ E" using CongA_def Q3 by blast ultimately show ?thesis using l11_10 out_trivial by blast qed then show ?thesis using not_conga_sym by blast qed then show ?thesis using P1 not_conga by blast qed ultimately show ?thesis using SumA_def by blast qed next assume W3: "\ Col D E F" then obtain K where W4: "A B C CongA F E K \ F E TS K D" using Q1 ex_conga_ts not_col_permutation_3 by blast show ?thesis proof - have "F E K CongA A B C" using W4 not_conga_sym by blast moreover have "\ E F OS D K" proof - have "E F TS D K" using W4 invert_two_sides l9_2 by blast then show ?thesis using l9_9 by auto qed moreover have "Coplanar D E F K" proof - have "E F TS D K" using W4 invert_two_sides l9_2 by blast then show ?thesis using ncoplanar_perm_8 ts__coplanar by blast qed moreover have "D E K CongA G H I" proof - have "A B J CongA K E D" proof - have "B C TS A J" proof - have "Coplanar B C A J" using P1 ncoplanar_perm_12 by blast moreover have "\ Col A B C" by (simp add: Q1) moreover have "\ B C OS A J" using P1 by simp moreover have "\ Col J B C" proof - { assume "Col J B C" have "Col D E F" proof - have "Col C B J" using Col_perm \Col J B C\ by blast moreover have "C B J CongA D E F" by (simp add: P1) ultimately show ?thesis using col_conga_col by blast qed then have "False" by (simp add: W3) } then show ?thesis by blast qed ultimately show ?thesis using cop_nos__ts by blast qed moreover have "E F TS K D" using W4 invert_two_sides by blast moreover have "A B C CongA K E F" by (simp add: W4 conga_right_comm) moreover have "C B J CongA F E D" by (simp add: P1 conga_right_comm) ultimately show ?thesis using l11_22a by auto qed then have "D E K CongA A B J" using conga_left_comm not_conga_sym by blast then show ?thesis using P1 not_conga by blast qed ultimately show ?thesis using SumA_def by blast qed qed qed qed lemma conga3_suma__suma: assumes "A B C D E F SumA G H I" and "A B C CongA A' B' C'" and "D E F CongA D' E' F'" and "G H I CongA G' H' I'" shows "A' B' C' D' E' F' SumA G' H' I'" proof - have "D' E' F' A B C SumA G' H' I'" proof - obtain J where P1: "C B J CongA D E F \ \ B C OS A J \ Coplanar A B C J \ A B J CongA G H I" using SumA_def assms(1) by blast have "A B C D' E' F' SumA G' H' I'" proof - have "C B J CongA D' E' F'" using P1 assms(3) not_conga by blast moreover have "\ B C OS A J" using P1 by simp moreover have "Coplanar A B C J" using P1 by simp moreover have "A B J CongA G' H' I'" using P1 assms(4) not_conga by blast ultimately show ?thesis using SumA_def by blast qed then show ?thesis by (simp add: suma_sym) qed then obtain J where P2: "F' E' J CongA A B C \ \ E' F' OS D' J \ Coplanar D' E' F' J \ D' E' J CongA G' H' I'" using SumA_def by blast have "D' E' F' A' B' C' SumA G' H' I'" proof - have "F' E' J CongA A' B' C'" proof - have "F' E' J CongA A B C" by (simp add: P2) moreover have "A B C CongA A' B' C'" by (simp add: assms(2)) ultimately show ?thesis using not_conga by blast qed moreover have "\ E' F' OS D' J" using P2 by simp moreover have "Coplanar D' E' F' J" using P2 by simp moreover have "D' E' J CongA G' H' I'" by (simp add: P2) ultimately show ?thesis using SumA_def by blast qed then show ?thesis by (simp add: suma_sym) qed lemma out6_suma__suma: assumes "A B C D E F SumA G H I" and "B Out A A'" and "B Out C C'" and "E Out D D'" and "E Out F F'" and "H Out G G'" and "H Out I I'" shows "A' B C' D' E F' SumA G' H I'" proof - have "A B C CongA A' B C'" using Out_cases assms(2) assms(3) out2__conga by blast moreover have "D E F CongA D' E F'" using Out_cases assms(4) assms(5) out2__conga by blast moreover have "G H I CongA G' H I'" by (simp add: assms(6) assms(7) l6_6 out2__conga) ultimately show ?thesis using assms(1) conga3_suma__suma by blast qed lemma out546_suma__conga: assumes "A B C D E F SumA G H I" and "E Out D F" shows "A B C CongA G H I" proof - have "A B C D E F SumA A B C" proof - have "C B C CongA D E F" by (metis assms(1) assms(2) l11_21_b out_trivial suma_distincts) moreover have "\ B C OS A C" using os_distincts by auto moreover have "Coplanar A B C C" using ncop_distincts by auto moreover have "A B C CongA A B C" by (metis Tarski_neutral_dimensionless.suma_distincts Tarski_neutral_dimensionless_axioms assms(1) conga_pseudo_refl conga_right_comm) ultimately show ?thesis using SumA_def by blast qed then show ?thesis using suma2__conga assms(1) by blast qed lemma out546__suma: assumes "A \ B" and "B \ C" and "E Out D F" shows "A B C D E F SumA A B C" proof - have P1: "D \ E" using assms(3) out_diff1 by auto have P2: "F \ E" using Out_def assms(3) by auto then obtain G H I where P3: "A B C D E F SumA G H I" using P1 assms(1) assms(2) ex_suma by presburger then have "G H I CongA A B C" by (meson Tarski_neutral_dimensionless.conga_sym Tarski_neutral_dimensionless.out546_suma__conga Tarski_neutral_dimensionless_axioms assms(3)) then show ?thesis using P1 P2 P3 assms(1) assms(2) assms(3) conga3_suma__suma conga_refl out_diff1 by auto qed lemma out213_suma__conga: assumes "A B C D E F SumA G H I" and "B Out A C" shows "D E F CongA G H I" using assms(1) assms(2) out546_suma__conga suma_sym by blast lemma out213__suma: assumes "D \ E" and "E \ F" and "B Out A C" shows "A B C D E F SumA D E F" by (simp add: assms(1) assms(2) assms(3) out546__suma suma_sym) lemma suma_left_comm: assumes "A B C D E F SumA G H I" shows "C B A D E F SumA G H I" proof - have "A B C CongA C B A" using assms conga_pseudo_refl suma_distincts by fastforce moreover have "D E F CongA D E F" by (metis assms conga_refl suma_distincts) moreover have "G H I CongA G H I" by (metis assms conga_refl suma_distincts) ultimately show ?thesis using assms conga3_suma__suma by blast qed lemma suma_middle_comm: assumes "A B C D E F SumA G H I" shows "A B C F E D SumA G H I" using assms suma_left_comm suma_sym by blast lemma suma_right_comm: assumes "A B C D E F SumA G H I" shows "A B C D E F SumA I H G" proof - have "A B C CongA A B C" using assms conga_refl suma_distincts by fastforce moreover have "D E F CongA D E F" by (metis assms conga_refl suma_distincts) moreover have "G H I CongA I H G" by (meson Tarski_neutral_dimensionless.conga_right_comm Tarski_neutral_dimensionless.suma2__conga Tarski_neutral_dimensionless_axioms assms) ultimately show ?thesis using assms conga3_suma__suma by blast qed lemma suma_comm: assumes "A B C D E F SumA G H I" shows "C B A F E D SumA I H G" by (simp add: assms suma_left_comm suma_middle_comm suma_right_comm) lemma ts__suma: assumes "A B TS C D" shows "C B A A B D SumA C B D" proof - have "A B D CongA A B D" by (metis Tarski_neutral_dimensionless.conga_right_comm Tarski_neutral_dimensionless_axioms assms conga_pseudo_refl ts_distincts) moreover have "\ B A OS C D" using assms invert_one_side l9_9 by blast moreover have "Coplanar C B A D" using assms ncoplanar_perm_14 ts__coplanar by blast moreover have "C B D CongA C B D" by (metis assms conga_refl ts_distincts) ultimately show ?thesis using SumA_def by blast qed lemma ts__suma_1: assumes "A B TS C D" shows "C A B B A D SumA C A D" by (simp add: assms invert_two_sides ts__suma) lemma inangle__suma: assumes "P InAngle A B C" shows "A B P P B C SumA A B C" proof - have "Coplanar A B P C" by (simp add: assms coplanar_perm_8 inangle__coplanar) moreover have "\ B P OS A C" by (meson assms col123__nos col124__nos in_angle_two_sides invert_two_sides l9_9_bis not_col_permutation_5) ultimately show ?thesis using SumA_def assms conga_refl inangle_distincts by blast qed lemma bet__suma: assumes "A \ B" and "B \ C" and "P \ B" and "Bet A B C" shows "A B P P B C SumA A B C" proof - have "P InAngle A B C" using assms(1) assms(2) assms(3) assms(4) in_angle_line by auto then show ?thesis by (simp add: inangle__suma) qed lemma sams_chara: assumes "A \ B" and "A' \ B" and "Bet A B A'" shows "SAMS A B C D E F \ D E F LeA C B A'" proof - { assume T1: "SAMS A B C D E F" obtain J where T2: "C B J CongA D E F \ \ B C OS A J \ \ A B TS C J \ Coplanar A B C J" using SAMS_def T1 by auto have T3: "A \ A'" using assms(2) assms(3) between_identity by blast have T4: "C \ B" using T2 conga_distinct by blast have T5: "J \ B" using T2 conga_diff2 by blast have T6: "D \ E" using CongA_def T2 by auto have T7: "F \ E" using CongA_def T2 by blast { assume "E Out D F" then have "D E F LeA C B A'" by (simp add: T4 assms(2) l11_31_1) } { assume T8: "\ Bet A B C" have "D E F LeA C B A'" proof cases assume "Col A B C" then have "Bet C B A'" using T8 assms(1) assms(3) between_exchange3 outer_transitivity_between2 third_point by blast then show ?thesis by (simp add: T4 T6 T7 assms(2) l11_31_2) next assume T9: "\ Col A B C" show ?thesis proof cases assume T10: "Col D E F" show ?thesis proof cases assume T11: "Bet D E F" have "D E F CongA C B J" by (simp add: T2 conga_sym) then have T12: "Bet C B J" using T11 bet_conga__bet by blast have "A B TS C J" proof - have "\ Col J A B" using T5 T9 T12 bet_col col2__eq col_permutation_1 by blast moreover have "\ T. Col T A B \ Bet C T J" using T12 col_trivial_3 by blast ultimately show ?thesis using T9 TS_def col_permutation_1 by blast qed then have "False" using T2 by simp then show ?thesis by simp next assume "\ Bet D E F" then show ?thesis using T10 \E Out D F \ D E F LeA C B A'\ or_bet_out by auto qed next assume T13: "\ Col D E F" show ?thesis proof - have "C B J LeA C B A'" proof - have "J InAngle C B A'" proof - have "A' \ B" by (simp add: assms(2)) moreover have "Bet A B A'" by (simp add: assms(3)) moreover have "C InAngle A B J" proof - have "\ Col J B C" proof - have "\ Col D E F" by (simp add: T13) moreover have "D E F CongA J B C" using T2 conga_left_comm not_conga_sym by blast ultimately show ?thesis using ncol_conga_ncol by blast qed then have "B C TS A J" by (simp add: T2 T9 cop_nos__ts coplanar_perm_8) then obtain X where T14: "Col X B C \ Bet A X J" using TS_def by blast { assume T15: "X \ B" have "B Out X C" proof - have "Col B X C" by (simp add: Col_perm T14) moreover have "B A OS X C" proof - have "A B OS X C" proof - have "A B OS X J" by (smt T14 T9 T15 bet_out calculation col_transitivity_2 col_trivial_2 l6_21 out_one_side) moreover have "A B OS J C" by (metis T14 T2 T9 calculation cop_nts__os l5_2 not_col_permutation_2 one_side_chara one_side_symmetry) ultimately show ?thesis using one_side_transitivity by blast qed then show ?thesis by (simp add: invert_one_side) qed ultimately show ?thesis using col_one_side_out by auto qed } then have "Bet A X J \ (X = B \ B Out X C)" using T14 by blast then show ?thesis using InAngle_def T4 T5 assms(1) by auto qed ultimately show ?thesis using in_angle_reverse l11_24 by blast qed moreover have "C B J CongA C B J" by (simp add: T4 T5 conga_refl) ultimately show ?thesis by (simp add: inangle__lea) qed moreover have "D E F LeA C B J" by (simp add: T2 conga__lea456123) ultimately show ?thesis using lea_trans by blast qed qed qed } then have "D E F LeA C B A'" using SAMS_def T1 \E Out D F \ D E F LeA C B A'\ by blast } { assume P1: "D E F LeA C B A'" have P2: "A \ A'" using assms(2) assms(3) between_identity by blast have P3: "C \ B" using P1 lea_distincts by auto have P4: "D \ E" using P1 lea_distincts by auto have P5: "F \ E" using P1 lea_distincts by auto have "SAMS A B C D E F" proof cases assume P6: "Col A B C" show ?thesis proof cases assume P7: "Bet A B C" have "E Out D F" proof - have "B Out C A'" by (meson Bet_perm P3 P7 assms(1) assms(2) assms(3) l6_2) moreover have "C B A' CongA D E F" using P1 calculation l11_21_b out_lea__out by blast ultimately show ?thesis using out_conga_out by blast qed moreover have "C B C CongA D E F" using P3 calculation l11_21_b out_trivial by auto moreover have "\ B C OS A C" using os_distincts by auto moreover have "\ A B TS C C" by (simp add: not_two_sides_id) moreover have "Coplanar A B C C" using ncop_distincts by auto ultimately show ?thesis using SAMS_def assms(1) by blast next assume P8: "\ Bet A B C" have P9: "B Out A C" by (simp add: P6 P8 l6_4_2) obtain J where P10: "D E F CongA C B J" using P3 P4 P5 angle_construction_3 by blast show ?thesis proof - have "C B J CongA D E F" using P10 not_conga_sym by blast moreover have "\ B C OS A J" using Col_cases P6 one_side_not_col123 by blast moreover have "\ A B TS C J" using Col_cases P6 TS_def by blast moreover have "Coplanar A B C J" using P6 col__coplanar by auto ultimately show ?thesis using P8 SAMS_def assms(1) by blast qed qed next assume P11: "\ Col A B C" have P12: "\ Col A' B C" using P11 assms(2) assms(3) bet_col bet_col1 colx by blast show ?thesis proof cases assume P13: "Col D E F" have P14: "E Out D F" proof - { assume P14: "Bet D E F" have "D E F LeA C B A'" by (simp add: P1) then have "Bet C B A'" using P14 bet_lea__bet by blast then have "Col A' B C" using Col_def Col_perm by blast then have "False" by (simp add: P12) } then have "\ Bet D E F" by auto then show ?thesis by (simp add: P13 l6_4_2) qed show ?thesis proof - have "C B C CongA D E F" by (simp add: P3 P14 l11_21_b out_trivial) moreover have "\ B C OS A C" using os_distincts by auto moreover have "\ A B TS C C" by (simp add: not_two_sides_id) moreover have "Coplanar A B C C" using ncop_distincts by auto ultimately show ?thesis using P14 SAMS_def assms(1) by blast qed next assume P15: "\ Col D E F" obtain J where P16: "D E F CongA C B J \ C B TS J A" using P11 P15 ex_conga_ts not_col_permutation_3 by presburger show ?thesis proof - have "C B J CongA D E F" by (simp add: P16 conga_sym) moreover have "\ B C OS A J" proof - have "C B TS A J" using P16 by (simp add: l9_2) then show ?thesis using invert_one_side l9_9 by blast qed moreover have "\ A B TS C J \ Coplanar A B C J" proof cases assume "Col A B J" then show ?thesis using TS_def ncop__ncols not_col_permutation_1 by blast next assume P17: "\ Col A B J" have "\ A B TS C J" proof - have "A' B OS J C" proof - have "\ Col A' B C" by (simp add: P12) moreover have "\ Col B A' J" proof - { assume "Col B A' J" then have "False" by (metis P17 assms(2) assms(3) bet_col col_trivial_2 colx) } then show ?thesis by auto qed moreover have "J InAngle A' B C" proof - obtain K where P20: "K InAngle C B A' \ D E F CongA C B K" using LeA_def P1 by blast have "J InAngle C B A'" proof - have "C B A' CongA C B A'" by (simp add: P3 assms(2) conga_pseudo_refl conga_right_comm) moreover have "C B K CongA C B J" proof - have "C B K CongA D E F" using P20 not_conga_sym by blast moreover have "D E F CongA C B J" by (simp add: P16) ultimately show ?thesis using not_conga by blast qed moreover have "K InAngle C B A'" using P20 by simp moreover have "C B OS J A'" proof - have "C B TS J A" using P16 by simp moreover have "C B TS A' A" using Col_perm P12 assms(3) bet__ts between_symmetry calculation invert_two_sides ts_distincts by blast ultimately show ?thesis using OS_def by auto qed ultimately show ?thesis using conga_preserves_in_angle by blast qed then show ?thesis by (simp add: l11_24) qed ultimately show ?thesis by (simp add: in_angle_one_side) qed then have "A' B OS C J" by (simp add: one_side_symmetry) then have "\ A' B TS C J" by (simp add: l9_9_bis) then show ?thesis using assms(2) assms(3) bet_col bet_col1 col_preserves_two_sides by blast qed moreover have "Coplanar A B C J" proof - have "C B TS J A" using P16 by simp then show ?thesis by (simp add: coplanar_perm_20 ts__coplanar) qed ultimately show ?thesis by auto qed ultimately show ?thesis using P11 SAMS_def assms(1) bet_col by auto qed qed qed } then show ?thesis using \SAMS A B C D E F \ D E F LeA C B A'\ by blast qed lemma sams_distincts: assumes "SAMS A B C D E F" shows "A \ B \ B \ C \ D \ E \ E \ F" proof - obtain J where P1: "C B J CongA D E F \ \ B C OS A J \ \ A B TS C J \ Coplanar A B C J" using SAMS_def assms by auto then show ?thesis by (metis SAMS_def assms conga_distinct) qed lemma sams_sym: assumes "SAMS A B C D E F" shows "SAMS D E F A B C" proof - have P1: "A \ B" using assms sams_distincts by blast have P3: "D \ E" using assms sams_distincts by blast obtain D' where P5: "E Midpoint D D'" using symmetric_point_construction by blast obtain A' where P6: "B Midpoint A A'" using symmetric_point_construction by blast have P8: "E \ D'" using P3 P5 is_midpoint_id_2 by blast have P9: "A \ A'" using P1 P6 l7_3 by auto then have P10: "B \ A'" using P6 P9 midpoint_not_midpoint by auto then have "D E F LeA C B A'" using P1 P6 assms midpoint_bet sams_chara by fastforce then have "D E F LeA A' B C" by (simp add: lea_right_comm) then have "A B C LeA D' E F" by (metis Mid_cases P1 P10 P3 P5 P6 P8 l11_36 midpoint_bet) then have "A B C LeA F E D'" by (simp add: lea_right_comm) moreover have "D \ E" by (simp add: P3) moreover have "D' \ E" using P8 by auto moreover have "Bet D E D'" by (simp add: P5 midpoint_bet) then show ?thesis using P3 P8 calculation(1) sams_chara by fastforce qed lemma sams_right_comm: assumes "SAMS A B C D E F" shows "SAMS A B C F E D" proof - have P1: "E Out D F \ \ Bet A B C" using SAMS_def assms by blast obtain J where P2: "C B J CongA D E F \ \ B C OS A J \ \ A B TS C J \ Coplanar A B C J" using SAMS_def assms by auto { assume "E Out D F" then have "E Out F D \ \ Bet A B C" by (simp add: l6_6) } { assume "\ Bet A B C" then have "E Out F D \ \ Bet A B C" by auto } then have "E Out F D \ \ Bet A B C" using \E Out D F \ E Out F D \ \ Bet A B C\ P1 by auto moreover have "C B J CongA F E D" proof - have "C B J CongA D E F" by (simp add: P2) then show ?thesis by (simp add: conga_right_comm) qed ultimately show ?thesis using P2 SAMS_def assms by auto qed lemma sams_left_comm: assumes "SAMS A B C D E F" shows "SAMS C B A D E F" proof - have "SAMS D E F A B C" by (simp add: assms sams_sym) then have "SAMS D E F C B A" using sams_right_comm by blast then show ?thesis using sams_sym by blast qed lemma sams_comm: assumes "SAMS A B C D E F" shows "SAMS C B A F E D" using assms sams_left_comm sams_right_comm by blast lemma conga2_sams__sams: assumes "A B C CongA A' B' C'" and "D E F CongA D' E' F'" and "SAMS A B C D E F" shows "SAMS A' B' C' D' E' F'" proof - obtain A0 where "B Midpoint A A0" using symmetric_point_construction by auto obtain A'0 where "B' Midpoint A' A'0" using symmetric_point_construction by blast have "D' E' F' LeA C' B' A'0" proof - have "D E F LeA C B A0" by (metis \B Midpoint A A0\ assms(1) assms(3) conga_distinct midpoint_bet midpoint_distinct_2 sams_chara) moreover have "D E F CongA D' E' F'" by (simp add: assms(2)) moreover have "C B A0 CongA C' B' A'0" proof - have "A0 B C CongA A'0 B' C'" by (metis \B Midpoint A A0\ \B' Midpoint A' A'0\ assms(1) calculation(1) conga_diff45 l11_13 lea_distincts midpoint_bet midpoint_not_midpoint) then show ?thesis using conga_comm by blast qed ultimately show ?thesis using l11_30 by blast qed then show ?thesis by (metis \B' Midpoint A' A'0\ assms(1) conga_distinct lea_distincts midpoint_bet sams_chara) qed lemma out546__sams: assumes "A \ B" and "B \ C" and "E Out D F" shows "SAMS A B C D E F" proof - obtain A' where "Bet A B A' \ Cong B A' A B" using segment_construction by blast then have "D E F LeA C B A'" using assms(1) assms(2) assms(3) cong_diff_3 l11_31_1 by fastforce then show ?thesis using \Bet A B A' \ Cong B A' A B\ assms(1) lea_distincts sams_chara by blast qed lemma out213__sams: assumes "D \ E" and "E \ F" and "B Out A C" shows "SAMS A B C D E F" by (simp add: Tarski_neutral_dimensionless.sams_sym Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) out546__sams) lemma bet_suma__sams: assumes "A B C D E F SumA G H I" and "Bet G H I" shows "SAMS A B C D E F" proof - obtain A' where P1: "C B A' CongA D E F \ \ B C OS A A' \ Coplanar A B C A' \ A B A' CongA G H I" using SumA_def assms(1) by auto then have "G H I CongA A B A'" using not_conga_sym by blast then have "Bet A B A'" using assms(2) bet_conga__bet by blast show ?thesis proof - have "E Out D F \ \ Bet A B C" proof - { assume "Bet A B C" then have "E Out D F" proof - have "B Out C A'" proof - have "C \ B" using assms(1) suma_distincts by blast moreover have "A' \ B" using CongA_def \G H I CongA A B A'\ by blast moreover have "A \ B" using CongA_def \G H I CongA A B A'\ by blast moreover have "Bet C B A" by (simp add: Bet_perm \Bet A B C\) ultimately show ?thesis using Out_def \Bet A B A'\ \Bet A B C\ l5_2 by auto qed moreover have "C B A' CongA D E F" using P1 by simp ultimately show ?thesis using l11_21_a by blast qed } then show ?thesis by blast qed moreover have "\ J. (C B J CongA D E F \ \ B C OS A J \ \ A B TS C J \ Coplanar A B C J)" proof - have "C B A' CongA D E F" by (simp add: P1) moreover have "\ B C OS A A'" by (simp add: P1) moreover have "\ A B TS C A'" using Col_def TS_def \Bet A B A'\ by blast moreover have "Coplanar A B C A'" by (simp add: P1) ultimately show ?thesis by blast qed ultimately show ?thesis using CongA_def SAMS_def \C B A' CongA D E F \ \ B C OS A A' \ Coplanar A B C A' \ A B A' CongA G H I\ by auto qed qed lemma bet__sams: assumes "A \ B" and "B \ C" and "P \ B" and "Bet A B C" shows "SAMS A B P P B C" by (meson assms(1) assms(2) assms(3) assms(4) bet__suma bet_suma__sams) lemma suppa__sams: assumes "A B C SuppA D E F" shows "SAMS A B C D E F" proof - obtain A' where P1: "Bet A B A' \ D E F CongA C B A'" using SuppA_def assms by auto then have "SAMS A B C C B A'" by (metis assms bet__sams conga_diff45 conga_diff56 suppa2__conga123) thus ?thesis by (meson P1 assms conga2_sams__sams not_conga_sym suppa2__conga123) qed lemma os_ts__sams: assumes "B P TS A C" and "A B OS P C" shows "SAMS A B P P B C" proof - have "B Out P C \ \ Bet A B P" using assms(2) bet_col col123__nos by blast moreover have "\ J. (P B J CongA P B C \ \ B P OS A J \ \ A B TS P J \ Coplanar A B P J)" by (metis assms(1) assms(2) conga_refl l9_9 os__coplanar os_distincts) ultimately show ?thesis using SAMS_def assms(2) os_distincts by auto qed lemma os2__sams: assumes "A B OS P C" and "C B OS P A" shows "SAMS A B P P B C" by (simp add: Tarski_neutral_dimensionless.os_ts__sams Tarski_neutral_dimensionless_axioms assms(1) assms(2) invert_one_side l9_31) lemma inangle__sams: assumes "P InAngle A B C" shows "SAMS A B P P B C" proof - have "Bet A B C \ B Out A C \ \ Col A B C" using l6_4_2 by blast { assume "Bet A B C" then have "SAMS A B P P B C" using assms bet__sams inangle_distincts by fastforce } { assume "B Out A C" then have "SAMS A B P P B C" by (metis assms in_angle_out inangle_distincts out213__sams) } { assume "\ Col A B C" then have "\ Bet A B C" using Col_def by auto { assume "Col B A P" have "SAMS A B P P B C" by (metis \Col B A P\ \\ Bet A B C\ assms col_in_angle_out inangle_distincts out213__sams) } { assume "\ Col B A P" { assume "Col B C P" have "SAMS A B P P B C" by (metis Tarski_neutral_dimensionless.sams_comm Tarski_neutral_dimensionless_axioms \Col B C P\ \\ Bet A B C\ assms between_symmetry col_in_angle_out inangle_distincts l11_24 out546__sams) } { assume "\ Col B C P" have "SAMS A B P P B C" proof - have "B P TS A C" by (simp add: \\ Col B A P\ \\ Col B C P\ assms in_angle_two_sides invert_two_sides) moreover have "A B OS P C" by (simp add: \\ Col A B C\ \\ Col B A P\ assms in_angle_one_side) ultimately show ?thesis by (simp add: os_ts__sams) qed } then have "SAMS A B P P B C" using \Col B C P \ SAMS A B P P B C\ by blast } then have "SAMS A B P P B C" using \Col B A P \ SAMS A B P P B C\ by blast } thus ?thesis using \B Out A C \ SAMS A B P P B C\ \Bet A B C \ SAMS A B P P B C\ \Bet A B C \ B Out A C \ \ Col A B C\ by blast qed lemma sams_suma__lea123789: assumes "A B C D E F SumA G H I" and "SAMS A B C D E F" shows "A B C LeA G H I" proof cases assume "Col A B C" show ?thesis proof cases assume "Bet A B C" have P1: "(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))" using SAMS_def assms(2) by auto { assume "E Out D F" then have "A B C CongA G H I" using assms(1) out546_suma__conga by auto then have "A B C LeA G H I" by (simp add: conga__lea) } thus ?thesis using P1 \Bet A B C\ by blast next assume "\ Bet A B C" then have "B Out A C" using \Col A B C\ or_bet_out by auto thus ?thesis by (metis assms(1) l11_31_1 suma_distincts) qed next assume "\ Col A B C" show ?thesis proof cases assume "Col D E F" show ?thesis proof cases assume "Bet D E F" have "SAMS D E F A B C" using assms(2) sams_sym by auto then have "B Out A C" using SAMS_def \Bet D E F\ by blast thus ?thesis using l11_31_1 by (metis assms(1) suma_distincts) next assume "\ Bet D E F" have "A B C CongA G H I" proof - have "A B C D E F SumA G H I" by (simp add: assms(1)) moreover have "E Out D F" using \Col D E F\ \\ Bet D E F\ l6_4_2 by auto ultimately show ?thesis using out546_suma__conga by auto qed show ?thesis by (simp add: \A B C CongA G H I\ conga__lea) qed next assume "\ Col D E F" show ?thesis proof cases assume "Col G H I" show ?thesis proof cases assume "Bet G H I" thus ?thesis by (metis assms(1) l11_31_2 suma_distincts) next assume "\ Bet G H I" then have "H Out G I" by (simp add: \Col G H I\ l6_4_2) have "E Out D F \ \ Bet A B C" using \\ Col A B C\ bet_col by auto { assume "\ Bet A B C" then obtain J where P2: "C B J CongA D E F \ \ B C OS A J \ Coplanar A B C J \ A B J CongA G H I" using SumA_def assms(1) by blast have "G H I CongA A B J" using P2 not_conga_sym by blast then have "B Out A J" using \H Out G I\ out_conga_out by blast then have "B C OS A J" using Col_perm \\ Col A B C\ out_one_side by blast then have "False" using \C B J CongA D E F \ \ B C OS A J \ Coplanar A B C J \ A B J CongA G H I\ by linarith } then have "False" using Col_def \\ Col A B C\ by blast thus ?thesis by blast qed next assume "\ Col G H I" obtain J where P4: "C B J CongA D E F \ \ B C OS A J \ \ A B TS C J \ Coplanar A B C J" using SAMS_def assms(2) by auto { assume "Col J B C" have "J B C CongA D E F" by (simp add: P4 conga_left_comm) then have "Col D E F" using col_conga_col \Col J B C\ by blast then have "False" using \\ Col D E F\ by blast } then have "\ Col J B C" by blast have "A B J CongA G H I" proof - have "A B C D E F SumA A B J" proof - have "C B J CongA D E F" using P4 by simp moreover have "\ B C OS A J" by (simp add: P4) moreover have "Coplanar A B C J" by (simp add: P4) moreover have "A B J CongA A B J" by (metis \\ Col A B C\ \\ Col J B C\ col_trivial_1 conga_refl) ultimately show ?thesis using SumA_def by blast qed then show ?thesis using assms(1) suma2__conga by auto qed have "\ Col J B A" using \A B J CongA G H I\ \\ Col G H I\ col_conga_col not_col_permutation_3 by blast have "A B C LeA A B J" proof - have "C InAngle A B J" by (metis Col_perm P4 \\ Col A B C\ \\ Col J B A\ \\ Col J B C\ cop_nos__ts coplanar_perm_7 coplanar_perm_8 invert_two_sides l9_2 os_ts__inangle) moreover have "A B C CongA A B C" using calculation in_angle_asym inangle3123 inangle_distincts by auto ultimately show ?thesis using inangle__lea by blast qed thus ?thesis using \A B J CongA G H I\ conga__lea lea_trans by blast qed qed qed lemma sams_suma__lea456789: assumes "A B C D E F SumA G H I" and "SAMS A B C D E F" shows "D E F LeA G H I" proof - have "D E F A B C SumA G H I" by (simp add: assms(1) suma_sym) moreover have "SAMS D E F A B C" using assms(2) sams_sym by blast ultimately show ?thesis using sams_suma__lea123789 by auto qed lemma sams_lea2__sams: assumes "SAMS A' B' C' D' E' F'" and "A B C LeA A' B' C'" and "D E F LeA D' E' F'" shows "SAMS A B C D E F" proof - obtain A0 where "B Midpoint A A0" using symmetric_point_construction by auto obtain A'0 where "B' Midpoint A' A'0" using symmetric_point_construction by auto have "D E F LeA C B A0" proof - have "D' E' F' LeA C B A0" proof - have "D' E' F' LeA C' B' A'0" by (metis \B' Midpoint A' A'0\ assms(1) assms(2) lea_distincts midpoint_bet midpoint_distinct_2 sams_chara) moreover have "C' B' A'0 LeA C B A0" by (metis Mid_cases \B Midpoint A A0\ \B' Midpoint A' A'0\ assms(2) l11_36_aux2 l7_3_2 lea_comm lea_distincts midpoint_bet sym_preserve_diff) ultimately show ?thesis using lea_trans by blast qed moreover have "D E F LeA D' E' F'" using assms(3) by auto ultimately show ?thesis using \D' E' F' LeA C B A0\ assms(3) lea_trans by blast qed then show ?thesis by (metis \B Midpoint A A0\ assms(2) lea_distincts midpoint_bet sams_chara) qed lemma sams_lea456_suma2__lea: assumes "D E F LeA D' E' F'" and "SAMS A B C D' E' F'" and "A B C D E F SumA G H I" and "A B C D' E' F' SumA G' H' I'" shows "G H I LeA G' H' I'" proof cases assume "E' Out D' F'" have "G H I CongA G' H' I'" proof - have "G H I CongA A B C" proof - have "A B C D E F SumA G H I" by (simp add: assms(3)) moreover have "E Out D F" using \E' Out D' F'\ assms(1) out_lea__out by blast ultimately show ?thesis using conga_sym out546_suma__conga by blast qed moreover have "A B C CongA G' H' I'" using \E' Out D' F'\ assms(4) out546_suma__conga by blast ultimately show ?thesis using conga_trans by blast qed thus ?thesis by (simp add: conga__lea) next assume T1: "\ E' Out D' F'" show ?thesis proof cases assume T2: "Col A B C" have "E' Out D' F' \ \ Bet A B C" using assms(2) SAMS_def by simp { assume "\ Bet A B C" then have "B Out A C" by (simp add: T2 l6_4_2) have "G H I LeA G' H' I'" proof - have "D E F LeA D' E' F'" by (simp add: assms(1)) moreover have "D E F CongA G H I" using \B Out A C\ assms(3) out213_suma__conga by auto moreover have "D' E' F' CongA G' H' I'" using \B Out A C\ assms(4) out213_suma__conga by auto ultimately show ?thesis using l11_30 by blast qed } thus ?thesis using T1 \E' Out D' F' \ \ Bet A B C\ by auto next assume "\ Col A B C" show ?thesis proof cases assume "Col D' E' F'" have "SAMS D' E' F' A B C" by (simp add: assms(2) sams_sym) { assume "\ Bet D' E' F'" then have "G H I LeA G' H' I'" using not_bet_out T1 \Col D' E' F'\ by auto } thus ?thesis by (metis assms(2) assms(3) assms(4) bet_lea__bet l11_31_2 sams_suma__lea456789 suma_distincts) next assume "\ Col D' E' F'" show ?thesis proof cases assume "Col D E F" have "\ Bet D E F" using bet_lea__bet Col_def \\ Col D' E' F'\ assms(1) by blast thus ?thesis proof - have "A B C LeA G' H' I'" using assms(2) assms(4) sams_suma__lea123789 by auto moreover have "A B C CongA G H I" by (meson \Col D E F\ \\ Bet D E F\ assms(3) or_bet_out out213_suma__conga suma_sym) moreover have "G' H' I' CongA G' H' I'" proof - have "G' \ H'" using calculation(1) lea_distincts by blast moreover have "H' \ I'" using \A B C LeA G' H' I'\ lea_distincts by blast ultimately show ?thesis using conga_refl by auto qed ultimately show ?thesis using l11_30 by blast qed next assume "\ Col D E F" show ?thesis proof cases assume "Col G' H' I'" show ?thesis proof cases assume "Bet G' H' I'" show ?thesis proof - have "G \ H" using assms(3) suma_distincts by auto moreover have "I \ H" using assms(3) suma_distincts by blast moreover have "G' \ H'" using assms(4) suma_distincts by auto moreover have "I' \ H'" using assms(4) suma_distincts by blast ultimately show ?thesis by (simp add: \Bet G' H' I'\ l11_31_2) qed next assume "\ Bet G' H' I'" have "B Out A C" proof - have "H' Out G' I'" using \Col G' H' I'\ l6_4_2 by (simp add: \\ Bet G' H' I'\) moreover have "A B C LeA G' H' I'" using sams_suma__lea123789 using assms(2) assms(4) by auto ultimately show ?thesis using out_lea__out by blast qed then have "Col A B C" using Col_perm out_col by blast then have "False" using \\ Col A B C\ by auto thus ?thesis by blast qed next assume "\ Col G' H' I'" obtain F'1 where P5: "C B F'1 CongA D' E' F' \ \ B C OS A F'1 \ \ A B TS C F'1 \ Coplanar A B C F'1" using SAMS_def assms(2) by auto have P6: "D E F LeA C B F'1" proof - have "D E F CongA D E F" using \\ Col D E F\ conga_refl not_col_distincts by fastforce moreover have "D' E' F' CongA C B F'1" by (simp add: P5 conga_sym) ultimately show ?thesis using assms(1) l11_30 by blast qed then obtain F1 where P6: "F1 InAngle C B F'1 \ D E F CongA C B F1" using LeA_def by auto have "A B F'1 CongA G' H' I'" proof - have "A B C D' E' F' SumA A B F'1" proof - have "C B F'1 CongA D' E' F'" using P5 by blast moreover have "\ B C OS A F'1" using P5 by auto moreover have "Coplanar A B C F'1" by (simp add: P5) moreover have "A B F'1 CongA A B F'1" proof - have "A \ B" using \\ Col A B C\ col_trivial_1 by blast moreover have "B \ F'1" using P6 inangle_distincts by auto ultimately show ?thesis using conga_refl by auto qed ultimately show ?thesis using SumA_def by blast qed moreover have "A B C D' E' F' SumA G' H' I'" by (simp add: assms(4)) ultimately show ?thesis using suma2__conga by auto qed have "\ Col A B F'1" using \A B F'1 CongA G' H' I'\ \\ Col G' H' I'\ col_conga_col by blast have "\ Col C B F'1" proof - have "\ Col D' E' F'" by (simp add: \\ Col D' E' F'\) moreover have "D' E' F' CongA C B F'1" using P5 not_conga_sym by blast ultimately show ?thesis using ncol_conga_ncol by blast qed show ?thesis proof - have "A B F1 LeA A B F'1" proof - have "F1 InAngle A B F'1" proof - have "F1 InAngle F'1 B A" proof - have "F1 InAngle F'1 B C" by (simp add: P6 l11_24) moreover have "C InAngle F'1 B A" proof - have "B C TS A F'1" using Col_perm P5 \\ Col A B C\ \\ Col C B F'1\ cop_nos__ts ncoplanar_perm_12 by blast obtain X where "Col X B C \ Bet A X F'1" using TS_def \B C TS A F'1\ by auto have "Bet F'1 X A" using Bet_perm \Col X B C \ Bet A X F'1\ by blast moreover have "(X = B) \ (B Out X C)" proof - have "B A OS X C" proof - have "A B OS X F'1" by (metis \Col X B C \ Bet A X F'1\ \\ Col A B C\ \\ Col A B F'1\ bet_out_1 calculation out_one_side) moreover have "A B OS F'1 C" using Col_perm P5 \\ Col A B C\ \\ Col A B F'1\ cop_nos__ts one_side_symmetry by blast ultimately show ?thesis using invert_one_side one_side_transitivity by blast qed thus ?thesis using Col_cases \Col X B C \ Bet A X F'1\ col_one_side_out by blast qed ultimately show ?thesis by (metis InAngle_def \\ Col A B C\ \\ Col A B F'1\ not_col_distincts) qed ultimately show ?thesis using in_angle_trans by blast qed then show ?thesis using l11_24 by blast qed moreover have "A B F1 CongA A B F1" proof - have "A \ B" using \\ Col A B C\ col_trivial_1 by blast moreover have "B \ F1" using P6 conga_diff56 by blast ultimately show ?thesis using conga_refl by auto qed ultimately show ?thesis by (simp add: inangle__lea) qed moreover have "A B F1 CongA G H I" proof - have "A B C D E F SumA A B F1" proof - have "B C TS F1 A" proof - have "B C TS F'1 A" proof - have "B C TS A F'1" using Col_perm P5 \\ Col A B C\ \\ Col C B F'1\ cop_nos__ts ncoplanar_perm_12 by blast thus ?thesis using l9_2 by blast qed moreover have "B C OS F'1 F1" proof - have "\ Col C B F'1" by (simp add: \\ Col C B F'1\) moreover have "\ Col B C F1" proof - have "\ Col D E F" using \\ Col D E F\ by auto moreover have "D E F CongA C B F1" by (simp add: P6) ultimately show ?thesis using ncol_conga_ncol not_col_permutation_4 by blast qed moreover have "F1 InAngle C B F'1" using P6 by blast ultimately show ?thesis using in_angle_one_side invert_one_side one_side_symmetry by blast qed ultimately show ?thesis using l9_8_2 by blast qed thus ?thesis proof - have "C B F1 CongA D E F" using P6 not_conga_sym by blast moreover have "\ B C OS A F1" using \B C TS F1 A\ l9_9 one_side_symmetry by blast moreover have "Coplanar A B C F1" using \B C TS F1 A\ ncoplanar_perm_9 ts__coplanar by blast moreover have "A B F1 CongA A B F1" proof - have "A \ B" using \\ Col A B C\ col_trivial_1 by blast moreover have "B \ F1" using P6 conga_diff56 by blast ultimately show ?thesis using conga_refl by auto qed ultimately show ?thesis using SumA_def by blast qed qed moreover have "A B C D E F SumA G H I" by (simp add: assms(3)) ultimately show ?thesis using suma2__conga by auto qed ultimately show ?thesis using \A B F'1 CongA G' H' I'\ l11_30 by blast qed qed qed qed qed qed lemma sams_lea123_suma2__lea: assumes "A B C LeA A' B' C'" and "SAMS A' B' C' D E F" and "A B C D E F SumA G H I" and "A' B' C' D E F SumA G' H' I'" shows "G H I LeA G' H' I'" by (meson assms(1) assms(2) assms(3) assms(4) sams_lea456_suma2__lea sams_sym suma_sym) lemma sams_lea2_suma2__lea: assumes "A B C LeA A' B' C'" and "D E F LeA D' E' F'" and "SAMS A' B' C' D' E' F'" and "A B C D E F SumA G H I" and "A' B' C' D' E' F' SumA G' H' I'" shows "G H I LeA G' H' I'" proof - obtain G'' H'' I'' where "A B C D' E' F' SumA G'' H'' I''" using assms(4) assms(5) ex_suma suma_distincts by presburger have "G H I LeA G'' H'' I''" proof - have "D E F LeA D' E' F'" by (simp add: assms(2)) moreover have "SAMS A B C D' E' F'" proof - have "SAMS A' B' C' D' E' F'" by (simp add: assms(3)) moreover have "A B C LeA A' B' C'" using assms(1) by auto moreover have "D' E' F' LeA D' E' F'" using assms(2) lea_distincts lea_refl by blast ultimately show ?thesis using sams_lea2__sams by blast qed moreover have "A B C D E F SumA G H I" by (simp add: assms(4)) moreover have "A B C D' E' F' SumA G'' H'' I''" by (simp add: \A B C D' E' F' SumA G'' H'' I''\) ultimately show ?thesis using sams_lea456_suma2__lea by blast qed moreover have "G'' H'' I'' LeA G' H' I'" using sams_lea123_suma2__lea assms(3) assms(5) \A B C D' E' F' SumA G'' H'' I''\ assms(1) by blast ultimately show ?thesis using lea_trans by blast qed lemma sams2_suma2__conga456: assumes "SAMS A B C D E F" and "SAMS A B C D' E' F'" and "A B C D E F SumA G H I" and "A B C D' E' F' SumA G H I" shows "D E F CongA D' E' F'" proof cases assume "Col A B C" show ?thesis proof cases assume P2: "Bet A B C" then have "E Out D F" using assms(1) SAMS_def by blast moreover have "E' Out D' F'" using P2 assms(2) SAMS_def by blast ultimately show ?thesis by (simp add: l11_21_b) next assume "\ Bet A B C" then have "B Out A C" using \Col A B C\ or_bet_out by blast show ?thesis proof - have "D E F CongA G H I" proof - have "A B C D E F SumA G H I" by (simp add: assms(3)) thus ?thesis using \B Out A C\ out213_suma__conga by auto qed moreover have "G H I CongA D' E' F'" proof - have "A B C D' E' F' SumA G H I" by (simp add: assms(4)) then have "D' E' F' CongA G H I" using \B Out A C\ out213_suma__conga by auto thus ?thesis using not_conga_sym by blast qed ultimately show ?thesis using not_conga by blast qed qed next assume "\ Col A B C" obtain J where T1: "C B J CongA D E F \ \ B C OS A J \ \ A B TS C J \ Coplanar A B C J" using assms(1) SAMS_def by blast have T1A: "C B J CongA D E F" using T1 by simp have T1B: "\ B C OS A J" using T1 by simp have T1C: "\ A B TS C J" using T1 by simp have T1D: "Coplanar A B C J" using T1 by simp obtain J' where T2: "C B J' CongA D' E' F' \ \ B C OS A J' \ \ A B TS C J' \ Coplanar A B C J'" using assms(2) SAMS_def by blast have T2A: "C B J' CongA D' E' F'" using T2 by simp have T2B: "\ B C OS A J'" using T2 by simp have T2C: "\ A B TS C J'" using T2 by simp have T2D: "Coplanar A B C J'" using T2 by simp have T3: "C B J CongA C B J'" proof - have "A B J CongA A B J'" proof - have "A B J CongA G H I" proof - have "A B C D E F SumA A B J" using SumA_def T1A T1B T1D \\ Col A B C\ conga_distinct conga_refl not_col_distincts by auto thus ?thesis using assms(3) suma2__conga by blast qed moreover have "G H I CongA A B J'" proof - have "A B C D' E' F' SumA G H I" by (simp add: assms(4)) moreover have "A B C D' E' F' SumA A B J'" using SumA_def T2A T2B T2D \\ Col A B C\ conga_distinct conga_refl not_col_distincts by auto ultimately show ?thesis using suma2__conga by auto qed ultimately show ?thesis using conga_trans by blast qed have "B Out J J' \ A B TS J J'" proof - have "Coplanar A B J J'" using T1D T2D \\ Col A B C\ coplanar_trans_1 ncoplanar_perm_8 not_col_permutation_2 by blast moreover have "A B J CongA A B J'" by (simp add: \A B J CongA A B J'\) ultimately show ?thesis by (simp add: conga_cop__or_out_ts) qed { assume "B Out J J'" then have "C B J CongA C B J'" by (metis Out_cases \\ Col A B C\ bet_out between_trivial not_col_distincts out2__conga) } { assume "A B TS J J'" then have "A B OS J C" by (meson T1C T1D TS_def \\ Col A B C\ cop_nts__os not_col_permutation_2 one_side_symmetry) then have "A B TS C J'" using \A B TS J J'\ l9_8_2 by blast then have "False" by (simp add: T2C) then have "C B J CongA C B J'" by blast } thus ?thesis using \B Out J J' \ C B J CongA C B J'\ \B Out J J' \ A B TS J J'\ by blast qed then have "C B J CongA D' E' F'" using T2A not_conga by blast thus ?thesis using T1A not_conga not_conga_sym by blast qed lemma sams2_suma2__conga123: assumes "SAMS A B C D E F" and "SAMS A' B' C' D E F" and "A B C D E F SumA G H I" and "A' B' C' D E F SumA G H I" shows "A B C CongA A' B' C'" proof - have "SAMS D E F A B C" by (simp add: assms(1) sams_sym) moreover have "SAMS D E F A' B' C'" by (simp add: assms(2) sams_sym) moreover have "D E F A B C SumA G H I" by (simp add: assms(3) suma_sym) moreover have "D E F A' B' C' SumA G H I" using assms(4) suma_sym by blast ultimately show ?thesis using sams2_suma2__conga456 by auto qed lemma suma_assoc_1: assumes "SAMS A B C D E F" and "SAMS D E F G H I" and "A B C D E F SumA A' B' C'" and "D E F G H I SumA D' E' F'" and "A' B' C' G H I SumA K L M" shows "A B C D' E' F' SumA K L M" proof - obtain A0 where P1: "Bet A B A0 \ Cong A B B A0" using Cong_perm segment_construction by blast obtain D0 where P2: "Bet D E D0 \ Cong D E E D0" using Cong_perm segment_construction by blast show ?thesis proof cases assume "Col A B C" show ?thesis proof cases assume "Bet A B C" then have "E Out D F" using SAMS_def assms(1) by simp show ?thesis proof - have "A' B' C' CongA A B C" using assms(3) \E Out D F\ conga_sym out546_suma__conga by blast moreover have "G H I CongA D' E' F'" using assms(4) \E Out D F\ out213_suma__conga by auto ultimately show ?thesis by (meson Tarski_neutral_dimensionless.conga3_suma__suma Tarski_neutral_dimensionless.suma2__conga Tarski_neutral_dimensionless_axioms assms(5)) qed next assume "\ Bet A B C" then have "B Out A C" using \Col A B C\ l6_4_2 by auto have "A \ B" using \B Out A C\ out_distinct by auto have "B \ C" using \\ Bet A B C\ between_trivial by auto have "D' \ E'" using assms(4) suma_distincts by blast have "E' \ F'" using assms(4) suma_distincts by auto show ?thesis proof - obtain K0 L0 M0 where P3:"A B C D' E' F' SumA K0 L0 M0" using ex_suma \A \ B\ \B \ C\ \D' \ E'\ \E' \ F'\ by presburger moreover have "A B C CongA A B C" using \A \ B\ \B \ C\ conga_refl by auto moreover have "D' E' F' CongA D' E' F'" using \D' \ E'\ \E' \ F'\ conga_refl by auto moreover have "K0 L0 M0 CongA K L M" proof - have "K0 L0 M0 CongA D' E' F'" using P3 \B Out A C\ conga_sym out213_suma__conga by blast moreover have "D' E' F' CongA K L M" proof - have "D E F G H I SumA D' E' F'" by (simp add: assms(4)) moreover have "D E F G H I SumA K L M" by (meson Tarski_neutral_dimensionless.conga3_suma__suma Tarski_neutral_dimensionless.out213_suma__conga Tarski_neutral_dimensionless.sams2_suma2__conga456 Tarski_neutral_dimensionless.suma2__conga Tarski_neutral_dimensionless_axioms \B Out A C\ assms(2) assms(3) assms(5) calculation not_conga_sym) ultimately show ?thesis using suma2__conga by auto qed ultimately show ?thesis using not_conga by blast qed ultimately show ?thesis using conga3_suma__suma by blast qed qed next assume T1: "\ Col A B C" have "\ Col C B A0" by (metis Col_def P1 \\ Col A B C\ cong_diff l6_16_1) show ?thesis proof cases assume "Col D E F" show ?thesis proof cases assume "Bet D E F" have "H Out G I" using SAMS_def assms(2) \Bet D E F\ by blast have "A B C D E F SumA A' B' C'" by (simp add: assms(3)) moreover have "A B C CongA A B C" by (metis \\ Col A B C\ conga_pseudo_refl conga_right_comm not_col_distincts) moreover have "D E F CongA D' E' F'" using \H Out G I\ assms(4) out546_suma__conga by auto moreover have "A' B' C' CongA K L M" using \H Out G I\ assms(5) out546_suma__conga by auto ultimately show ?thesis using conga3_suma__suma by blast next assume "\ Bet D E F" then have "E Out D F" using not_bet_out by (simp add: \Col D E F\) show ?thesis proof - have "A' B' C' CongA A B C" using assms(3) \E Out D F\ conga_sym out546_suma__conga by blast moreover have "G H I CongA D' E' F'" using out546_suma__conga \E Out D F\ assms(4) out213_suma__conga by auto moreover have "K L M CongA K L M" proof - have "K \ L \ L \ M" using assms(5) suma_distincts by blast thus ?thesis using conga_refl by auto qed ultimately show ?thesis using assms(5) conga3_suma__suma by blast qed qed next assume "\ Col D E F" then have "\ Col F E D0" by (metis Col_def P2 cong_diff l6_16_1 not_col_distincts) show ?thesis proof cases assume "Col G H I" show ?thesis proof cases assume "Bet G H I" have "SAMS G H I D E F" by (simp add: assms(2) sams_sym) then have "E Out D F" using SAMS_def \Bet G H I\ by blast then have "Col D E F" using Col_perm out_col by blast then have "False" using \\ Col D E F\ by auto thus ?thesis by simp next assume "\ Bet G H I" then have "H Out G I" using SAMS_def by (simp add: \Col G H I\ l6_4_2) show ?thesis proof - have "A B C CongA A B C" by (metis \\ Col A B C\ conga_refl not_col_distincts) moreover have "D E F CongA D' E' F'" using assms(4) out546_suma__conga \H Out G I\ by auto moreover have "A' B' C' CongA K L M" using \H Out G I\ assms(5) out546_suma__conga by auto ultimately show ?thesis using assms(3) conga3_suma__suma by blast qed qed next assume "\ Col G H I" have "\ B C OS A A0" using P1 col_trivial_1 one_side_chara by blast have "E F TS D D0" by (metis P2 \\ Col D E F\ \\ Col F E D0\ bet__ts bet_col between_trivial not_col_permutation_1) show ?thesis proof cases assume "Col A' B' C'" show ?thesis proof cases assume "Bet A' B' C'" show ?thesis proof cases assume "Col D' E' F'" show ?thesis proof cases assume "Bet D' E' F'" have "A B C CongA G H I" proof - have "A B C CongA D0 E F" proof - have "SAMS A B C D E F" by (simp add: assms(1)) moreover have "SAMS D0 E F D E F" by (metis P2 \\ Col D E F\ \\ Col F E D0\ bet__sams between_symmetry not_col_distincts sams_right_comm) moreover have "A B C D E F SumA A' B' C'" by (simp add: assms(3)) moreover have "D0 E F D E F SumA A' B' C'" proof - have "D E F D0 E F SumA A' B' C'" proof - have "F E D0 CongA D0 E F" by (metis \\ Col F E D0\ col_trivial_1 col_trivial_2 conga_pseudo_refl) moreover have "\ E F OS D D0" using P2 col_trivial_1 one_side_chara by blast moreover have "Coplanar D E F D0" by (meson P2 bet__coplanar ncoplanar_perm_1) moreover have "D E D0 CongA A' B' C'" using assms(3) P2 \Bet A' B' C'\ \\ Col F E D0\ conga_line not_col_distincts suma_distincts by auto ultimately show ?thesis using SumA_def by blast qed thus ?thesis by (simp add: \D E F D0 E F SumA A' B' C'\ suma_sym) qed ultimately show ?thesis using sams2_suma2__conga123 by blast qed moreover have "D0 E F CongA G H I" proof - have "SAMS D E F D0 E F" using P2 \\ Col D E F\ \\ Col F E D0\ bet__sams not_col_distincts sams_right_comm by auto moreover have "D E F D0 E F SumA D' E' F'" proof - have "F E D0 CongA D0 E F" by (metis (no_types) \\ Col F E D0\ col_trivial_1 col_trivial_2 conga_pseudo_refl) moreover have "\ E F OS D D0" using P2 col_trivial_1 one_side_chara by blast moreover have "Coplanar D E F D0" using P2 bet__coplanar ncoplanar_perm_1 by blast moreover have "D E D0 CongA D' E' F'" using assms(3) P2 \Bet D' E' F'\ \\ Col F E D0\ assms(4) conga_line not_col_distincts suma_distincts by auto ultimately show ?thesis using SumA_def by blast qed ultimately show ?thesis using assms(2) assms(4) sams2_suma2__conga456 by auto qed ultimately show ?thesis using conga_trans by blast qed then have "G H I CongA A B C" using not_conga_sym by blast have "D' E' F' A B C SumA K L M" proof - have "A' B' C' CongA D' E' F'" by (metis Tarski_neutral_dimensionless.suma_distincts Tarski_neutral_dimensionless_axioms \Bet A' B' C'\ \Bet D' E' F'\ assms(4) assms(5) conga_line) then show ?thesis by (meson Tarski_neutral_dimensionless.conga3_suma__suma Tarski_neutral_dimensionless.suma2__conga Tarski_neutral_dimensionless_axioms \G H I CongA A B C\ assms(5)) qed thus ?thesis by (simp add: suma_sym) next assume "\ Bet D' E' F'" then have "E' Out D' F'" by (simp add: \Col D' E' F'\ l6_4_2) have "D E F LeA D' E' F'" using assms(2) assms(4) sams_suma__lea123789 by blast then have "E Out D F" using \E' Out D' F'\ out_lea__out by blast then have "Col D E F" using Col_perm out_col by blast then have "False" using \\ Col D E F\ by auto thus ?thesis by simp qed next assume "\ Col D' E' F'" have "D E F CongA C B A0" proof - have "SAMS A B C D E F" by (simp add: assms(1)) moreover have "SAMS A B C C B A0" using P1 \\ Col A B C\ \\ Col C B A0\ bet__sams not_col_distincts by auto moreover have "A B C D E F SumA A' B' C'" by (simp add: assms(3)) moreover have "A B C C B A0 SumA A' B' C'" proof - have "A B C C B A0 SumA A B A0" by (metis P1 \\ Col A B C\ \\ Col C B A0\ bet__suma col_trivial_1 col_trivial_2) moreover have "A B C CongA A B C" using \SAMS A B C C B A0\ calculation sams2_suma2__conga123 by auto moreover have "C B A0 CongA C B A0" using \SAMS A B C C B A0\ calculation(1) sams2_suma2__conga456 by auto moreover have "A B A0 CongA A' B' C'" using P1 \Bet A' B' C'\ \\ Col C B A0\ assms(3) conga_line not_col_distincts suma_distincts by auto ultimately show ?thesis using conga3_suma__suma by blast qed ultimately show ?thesis using sams2_suma2__conga456 by blast qed have "SAMS C B A0 G H I" proof - have "D E F CongA C B A0" by (simp add: \D E F CongA C B A0\) moreover have "G H I CongA G H I" using \\ Col G H I\ conga_refl not_col_distincts by fastforce moreover have "SAMS D E F G H I" by (simp add: assms(2)) ultimately show ?thesis using conga2_sams__sams by blast qed then obtain J where P3: "A0 B J CongA G H I \ \ B A0 OS C J \ \ C B TS A0 J \ Coplanar C B A0 J" using SAMS_def by blast obtain F1 where P4: "F E F1 CongA G H I \ \ E F OS D F1 \ \ D E TS F F1 \ Coplanar D E F F1" using SAMS_def assms(2) by auto have "C B J CongA D' E' F'" proof - have "C B J CongA D E F1" proof - have "(B A0 TS C J \ E F TS D F1) \ (B A0 OS C J \ E F OS D F1)" proof - have "B A0 TS C J" proof - have "Coplanar B A0 C J" using P3 ncoplanar_perm_12 by blast moreover have "\ Col C B A0" by (simp add: \\ Col C B A0\) moreover have "\ Col J B A0" using P3 \\ Col G H I\ col_conga_col not_col_permutation_3 by blast moreover have "\ B A0 OS C J" using P3 by simp ultimately show ?thesis by (simp add: cop_nos__ts) qed moreover have "E F TS D F1" proof - have "Coplanar E F D F1" using P4 ncoplanar_perm_12 by blast moreover have "\ Col D E F" by (simp add: \\ Col D E F\) moreover have "\ Col F1 E F" using P4 \\ Col G H I\ col_conga_col col_permutation_3 by blast moreover have "\ E F OS D F1" using P4 by auto ultimately show ?thesis by (simp add: cop_nos__ts) qed ultimately show ?thesis by simp qed moreover have "C B A0 CongA D E F" using \D E F CongA C B A0\ not_conga_sym by blast moreover have "A0 B J CongA F E F1" proof - have "A0 B J CongA G H I" by (simp add: P3) moreover have "G H I CongA F E F1" using P4 not_conga_sym by blast ultimately show ?thesis using conga_trans by blast qed ultimately show ?thesis using l11_22 by auto qed moreover have "D E F1 CongA D' E' F'" proof - have "D E F G H I SumA D E F1" using P4 SumA_def \\ Col D E F\ conga_distinct conga_refl not_col_distincts by auto moreover have "D E F G H I SumA D' E' F'" by (simp add: assms(4)) ultimately show ?thesis using suma2__conga by auto qed ultimately show ?thesis using conga_trans by blast qed show ?thesis proof - have "A B C D' E' F' SumA A B J" proof - have "C B TS J A" proof - have "C B TS A0 A" proof - have "B \ A0" using \\ Col C B A0\ not_col_distincts by blast moreover have "\ Col B C A" using Col_cases \\ Col A B C\ by auto moreover have "Bet A B A0" by (simp add: P1) ultimately show ?thesis by (metis Bet_cases Col_cases \\ Col C B A0\ bet__ts invert_two_sides not_col_distincts) qed moreover have "C B OS A0 J" proof - have "\ Col J C B" using \C B J CongA D' E' F'\ \\ Col D' E' F'\ col_conga_col not_col_permutation_2 by blast moreover have "\ Col A0 C B" using Col_cases \\ Col C B A0\ by blast ultimately show ?thesis using P3 cop_nos__ts by blast qed ultimately show ?thesis using l9_8_2 by blast qed moreover have "C B J CongA D' E' F'" by (simp add: \C B J CongA D' E' F'\) moreover have "\ B C OS A J" using calculation(1) invert_one_side l9_9_bis one_side_symmetry by blast moreover have "Coplanar A B C J" using calculation(1) ncoplanar_perm_15 ts__coplanar by blast moreover have "A B J CongA A B J" proof - have "A \ B" using \\ Col A B C\ col_trivial_1 by auto moreover have "B \ J" using \C B TS J A\ ts_distincts by blast ultimately show ?thesis using conga_refl by auto qed ultimately show ?thesis using SumA_def by blast qed moreover have "A B J CongA K L M" proof - have "A' B' C' G H I SumA A B J" proof - have "A B A0 CongA A' B' C'" using P1 \Bet A' B' C'\ \\ Col A B C\ \\ Col C B A0\ assms(5) conga_line not_col_distincts suma_distincts by auto moreover have "A0 B J CongA G H I" by (simp add: P3) moreover have "A B A0 A0 B J SumA A B J" proof - have "A0 B J CongA A0 B J" proof - have "A0 \ B" using \\ Col C B A0\ col_trivial_2 by auto moreover have "B \ J" using CongA_def \A0 B J CongA G H I\ by blast ultimately show ?thesis using conga_refl by auto qed moreover have "\ B A0 OS A J" by (simp add: Col_def P1 col123__nos) moreover have "Coplanar A B A0 J" using P1 bet__coplanar by auto moreover have "A B J CongA A B J" using P1 \\ Col A B C\ between_symmetry calculation(1) l11_13 not_col_distincts by blast ultimately show ?thesis using SumA_def by blast qed ultimately show ?thesis by (meson conga3_suma__suma suma2__conga) qed moreover have "A' B' C' G H I SumA K L M" by (simp add: assms(5)) ultimately show ?thesis using suma2__conga by auto qed ultimately show ?thesis proof - have "A B C CongA A B C \ D' E' F' CongA D' E' F'" using CongA_def \A B J CongA K L M\ \C B J CongA D' E' F'\ conga_refl by presburger then show ?thesis using \A B C D' E' F' SumA A B J\ \A B J CongA K L M\ conga3_suma__suma by blast qed qed qed next assume "\ Bet A' B' C'" have "B Out A C" proof - have "A B C LeA A' B' C'" using assms(1) assms(3) sams_suma__lea123789 by auto moreover have "B' Out A' C'" using \Col A' B' C'\ \\ Bet A' B' C'\ or_bet_out by blast ultimately show ?thesis using out_lea__out by blast qed then have "Col A B C" using Col_perm out_col by blast then have "False" using \\ Col A B C\ by auto thus ?thesis by simp qed next assume "\ Col A' B' C'" obtain C1 where P6: "C B C1 CongA D E F \ \ B C OS A C1 \ \ A B TS C C1 \ Coplanar A B C C1" using SAMS_def assms(1) by auto have P6A: "C B C1 CongA D E F" using P6 by simp have P6B: "\ B C OS A C1" using P6 by simp have P6C: "\ A B TS C C1" using P6 by simp have P6D: "Coplanar A B C C1" using P6 by simp have "A' B' C' CongA A B C1" proof - have "A B C D E F SumA A B C1" using P6A P6B P6D SumA_def \\ Col A B C\ conga_distinct conga_refl not_col_distincts by auto moreover have "A B C D E F SumA A' B' C'" by (simp add: assms(3)) ultimately show ?thesis using suma2__conga by auto qed have "B C1 OS C A" proof - have "B A OS C C1" proof - have "A B OS C C1" proof - have "\ Col C A B" using Col_perm \\ Col A B C\ by blast moreover have "\ Col C1 A B" using \\ Col A' B' C'\ \A' B' C' CongA A B C1\ col_permutation_1 ncol_conga_ncol by blast ultimately show ?thesis using P6C P6D cop_nos__ts by blast qed thus ?thesis by (simp add: invert_one_side) qed moreover have "B C TS A C1" proof - have "Coplanar B C A C1" using P6D ncoplanar_perm_12 by blast moreover have "\ Col C1 B C" proof - have "D E F CongA C1 B C" using P6A conga_left_comm not_conga_sym by blast thus ?thesis using \\ Col D E F\ ncol_conga_ncol by blast qed ultimately show ?thesis using T1 P6B cop_nos__ts by blast qed ultimately show ?thesis using os_ts1324__os one_side_symmetry by blast qed show ?thesis proof cases assume "Col D' E' F'" show ?thesis proof cases assume "Bet D' E' F'" obtain C0 where P7: "Bet C B C0 \ Cong C B B C0" using Cong_perm segment_construction by blast have "B C1 TS C C0" by (metis P7 \B C1 OS C A\ bet__ts cong_diff_2 not_col_distincts one_side_not_col123) show ?thesis proof - have "A B C C B C0 SumA A B C0" proof - have "C B C0 CongA C B C0" by (metis P7 T1 cong_diff conga_line not_col_distincts) moreover have "\ B C OS A C0" using P7 bet_col col124__nos invert_one_side by blast moreover have "Coplanar A B C C0" using P7 bet__coplanar ncoplanar_perm_15 by blast moreover have "A B C0 CongA A B C0" by (metis P7 T1 cong_diff conga_refl not_col_distincts) ultimately show ?thesis using SumA_def by blast qed moreover have "A B C0 CongA K L M" proof - have "A' B' C' G H I SumA A B C0" proof - have "A B C1 C1 B C0 SumA A B C0" using \B C1 TS C C0\ \B C1 OS C A\ l9_8_2 ts__suma_1 by blast moreover have "A B C1 CongA A' B' C'" by (simp add: P6 \A' B' C' CongA A B C1\ conga_sym) moreover have "C1 B C0 CongA G H I" proof - have "SAMS C B C1 C1 B C0" by (metis P7 \B C1 TS C C0\ bet__sams ts_distincts) moreover have "SAMS C B C1 G H I" proof - have "D E F CongA C B C1" using P6A not_conga_sym by blast moreover have "G H I CongA G H I" using \\ Col G H I\ conga_refl not_col_distincts by fastforce moreover have "SAMS D E F G H I" by (simp add: assms(2)) ultimately show ?thesis using conga2_sams__sams by blast qed moreover have "C B C1 C1 B C0 SumA C B C0" by (simp add: \B C1 TS C C0\ ts__suma_1) moreover have "C B C1 G H I SumA C B C0" proof - have "D E F G H I SumA D' E' F'" by (simp add: assms(4)) moreover have "D E F CongA C B C1" using P6A not_conga_sym by blast moreover have "G H I CongA G H I" using \\ Col G H I\ conga_refl not_col_distincts by fastforce moreover have "D' E' F' CongA C B C0" using P7 assms(4) by (metis P6A Tarski_neutral_dimensionless.suma_distincts Tarski_neutral_dimensionless_axioms \Bet D' E' F'\ cong_diff conga_diff1 conga_line) ultimately show ?thesis using conga3_suma__suma by blast qed ultimately show ?thesis using sams2_suma2__conga456 by auto qed moreover have "A B C0 CongA A B C0" by (metis P7 T1 cong_diff conga_refl not_col_distincts) ultimately show ?thesis using conga3_suma__suma by blast qed thus ?thesis using assms(5) suma2__conga by auto qed moreover have "A B C CongA A B C" proof - have "A \ B \ B \ C" using T1 col_trivial_1 col_trivial_2 by auto thus ?thesis using conga_refl by auto qed moreover have "C B C0 CongA D' E' F'" proof - have "C \ B" using T1 col_trivial_2 by blast moreover have "B \ C0" using \B C1 TS C C0\ ts_distincts by blast moreover have "D' \ E'" using assms(4) suma_distincts by blast moreover have "E' \ F'" using assms(4) suma_distincts by auto ultimately show ?thesis by (simp add: P7 \Bet D' E' F'\ conga_line) qed ultimately show ?thesis using conga3_suma__suma by blast qed next assume "\ Bet D' E' F'" then have "E' Out D' F'" by (simp add: \Col D' E' F'\ l6_4_2) have "D E F LeA D' E' F'" using sams_suma__lea123789 assms(2) assms(4) by auto then have "E Out D F" using \E' Out D' F'\ out_lea__out by blast then have "False" using Col_cases \\ Col D E F\ out_col by blast thus ?thesis by simp qed next assume "\ Col D' E' F'" have "SAMS C B C1 G H I" proof - have "D E F CongA C B C1" using P6A not_conga_sym by blast moreover have "G H I CongA G H I" using \\ Col G H I\ conga_refl not_col_distincts by fastforce ultimately show ?thesis using assms(2) conga2_sams__sams by blast qed then obtain J where P7: "C1 B J CongA G H I \ \ B C1 OS C J \ \ C B TS C1 J \ Coplanar C B C1 J" using SAMS_def by blast have P7A: "C1 B J CongA G H I" using P7 by simp have P7B: "\ B C1 OS C J" using P7 by simp have P7C: "\ C B TS C1 J" using P7 by simp have P7D: "Coplanar C B C1 J" using P7 by simp obtain F1 where P8: "F E F1 CongA G H I \ \ E F OS D F1 \ \ D E TS F F1 \ Coplanar D E F F1" using SAMS_def assms(2) by auto have P8A: "F E F1 CongA G H I" using P8 by simp have P8B: "\ E F OS D F1" using P8 by simp have P8C: "\ D E TS F F1" using P8 by simp have P8D: "Coplanar D E F F1" using P8 by simp have "C B J CongA D' E' F'" proof - have "C B J CongA D E F1" proof - have "B C1 TS C J" proof - have "Coplanar B C1 C J" using P7D ncoplanar_perm_12 by blast moreover have "\ Col C B C1" using \B C1 OS C A\ not_col_permutation_2 one_side_not_col123 by blast moreover have "\ Col J B C1" using P7 \\ Col G H I\ col_conga_col not_col_permutation_3 by blast moreover have "\ B C1 OS C J" by (simp add: P7B) ultimately show ?thesis by (simp add: cop_nos__ts) qed moreover have "E F TS D F1" proof - have "Coplanar E F D F1" using P8D ncoplanar_perm_12 by blast moreover have "\ Col F1 E F" using P8 \\ Col G H I\ col_conga_col not_col_permutation_3 by blast ultimately show ?thesis using P8B \\ Col D E F\ cop_nos__ts by blast qed moreover have "C B C1 CongA D E F" using P6A by blast moreover have "C1 B J CongA F E F1" using P8 by (meson P7A not_conga not_conga_sym) ultimately show ?thesis using l11_22a by blast qed moreover have "D E F1 CongA D' E' F'" proof - have "D E F G H I SumA D E F1" using P8A P8B P8D SumA_def \\ Col D E F\ conga_distinct conga_refl not_col_distincts by auto moreover have "D E F G H I SumA D' E' F'" by (simp add: assms(4)) ultimately show ?thesis using suma2__conga by auto qed ultimately show ?thesis using conga_trans by blast qed have "\ Col C B C1" using \B C1 OS C A\ col123__nos col_permutation_1 by blast show ?thesis proof - have "A B C C B J SumA A B J" proof - have "B C TS J A" proof - have "B C TS C1 A" using cop_nos__ts using P6B P6D T1 \\ Col C B C1\ l9_2 ncoplanar_perm_12 not_col_permutation_3 by blast moreover have "B C OS C1 J" proof - have "\ Col C1 C B" using Col_perm \\ Col C B C1\ by blast moreover have "\ Col J C B" using \C B J CongA D' E' F'\ \\ Col D' E' F'\ col_conga_col col_permutation_1 by blast ultimately show ?thesis using P7C P7D cop_nos__ts invert_one_side by blast qed ultimately show ?thesis using l9_8_2 by blast qed thus ?thesis by (simp add: l9_2 ts__suma_1) qed moreover have "A B C CongA A B C" using T1 conga_refl not_col_distincts by fastforce moreover have "A B J CongA K L M" proof - have "A' B' C' G H I SumA A B J" proof - have "A B C1 C1 B J SumA A B J" proof - have "B C1 TS A J" proof - have "B C1 TS C J" proof - have "Coplanar B C1 C J" using P7D ncoplanar_perm_12 by blast moreover have "\ Col J B C1" using P7 \\ Col G H I\ col_conga_col not_col_permutation_3 by blast ultimately show ?thesis by (simp add: \\ Col C B C1\ P7B cop_nos__ts) qed moreover have "B C1 OS C A" by (simp add: \B C1 OS C A\) ultimately show ?thesis using l9_8_2 by blast qed thus ?thesis by (simp add: ts__suma_1) qed moreover have "A B C1 CongA A' B' C'" using \A' B' C' CongA A B C1\ not_conga_sym by blast moreover have "C1 B J CongA G H I" by (simp add: P7A) moreover have "A B J CongA A B J" using \A B C C B J SumA A B J\ suma2__conga by auto ultimately show ?thesis using conga3_suma__suma by blast qed moreover have "A' B' C' G H I SumA K L M" using assms(5) by simp ultimately show ?thesis using suma2__conga by auto qed ultimately show ?thesis using \C B J CongA D' E' F'\ conga3_suma__suma by blast qed qed qed qed qed qed qed lemma suma_assoc_2: assumes "SAMS A B C D E F" and "SAMS D E F G H I" and "A B C D E F SumA A' B' C'" and "D E F G H I SumA D' E' F'" and "A B C D' E' F' SumA K L M" shows "A' B' C' G H I SumA K L M" by (meson assms(1) assms(2) assms(3) assms(4) assms(5) sams_sym suma_assoc_1 suma_sym) lemma suma_assoc: assumes "SAMS A B C D E F" and "SAMS D E F G H I" and "A B C D E F SumA A' B' C'" and "D E F G H I SumA D' E' F'" shows "A' B' C' G H I SumA K L M \ A B C D' E' F' SumA K L M" by (meson assms(1) assms(2) assms(3) assms(4) suma_assoc_1 suma_assoc_2) lemma sams_assoc_1: assumes "SAMS A B C D E F" and "SAMS D E F G H I" and "A B C D E F SumA A' B' C'" and "D E F G H I SumA D' E' F'" and "SAMS A' B' C' G H I" shows "SAMS A B C D' E' F'" proof cases assume "Col A B C" { assume "E Out D F" have "SAMS A B C D' E' F'" proof - have "A' B' C' CongA A B C" using assms(3) \E Out D F\ conga_sym out546_suma__conga by blast moreover have "G H I CongA D' E' F'" using \E Out D F\ assms(4) out213_suma__conga by blast ultimately show ?thesis using assms(5) conga2_sams__sams by blast qed } { assume "\ Bet A B C" then have P1: "B Out A C" using \Col A B C\ or_bet_out by blast have "SAMS D' E' F' A B C" proof - have "D' \ E'" using assms(4) suma_distincts by auto moreover have "F' E' F' CongA A B C" proof - have "E' \ F'" using assms(4) suma_distincts by blast then have "E' Out F' F'" using out_trivial by auto thus ?thesis using P1 l11_21_b by blast qed moreover have "\ E' F' OS D' F'" using os_distincts by blast moreover have "\ D' E' TS F' F'" by (simp add: not_two_sides_id) moreover have "Coplanar D' E' F' F'" using ncop_distincts by blast ultimately show ?thesis using SAMS_def P1 by blast qed then have "SAMS A B C D' E' F'" using sams_sym by blast } thus ?thesis using SAMS_def assms(1) \E Out D F \ SAMS A B C D' E' F'\ by blast next assume "\ Col A B C" show ?thesis proof cases assume "Col D E F" have "H Out G I \ \ Bet D E F" using SAMS_def assms(2) by blast { assume "H Out G I" have "SAMS A B C D' E' F'" proof - have "A B C CongA A B C" using \\ Col A B C\ conga_refl not_col_distincts by fastforce moreover have "D E F CongA D' E' F'" using \H Out G I\ assms(4) out546_suma__conga by blast ultimately show ?thesis using assms(1) conga2_sams__sams by blast qed } { assume "\ Bet D E F" then have "E Out D F" using \Col D E F\ l6_4_2 by blast have "SAMS A B C D' E' F'" proof - have "A' B' C' CongA A B C" using out546_suma__conga \E Out D F\ assms(3) not_conga_sym by blast moreover have "G H I CongA D' E' F'" using out213_suma__conga \E Out D F\ assms(4) by auto ultimately show ?thesis using assms(5) conga2_sams__sams by auto qed } thus ?thesis using \H Out G I \ SAMS A B C D' E' F'\ \H Out G I \ \ Bet D E F\ by blast next assume "\ Col D E F" show ?thesis proof cases assume "Col G H I" have "SAMS G H I D E F" by (simp add: assms(2) sams_sym) then have "E Out D F \ \ Bet G H I" using SAMS_def by blast { assume "E Out D F" then have "False" using Col_cases \\ Col D E F\ out_col by blast then have "SAMS A B C D' E' F'" by simp } { assume "\ Bet G H I" then have "H Out G I" by (simp add: \Col G H I\ l6_4_2) have "SAMS A B C D' E' F'" proof - have "A B C CongA A B C" by (metis \\ Col A B C\ col_trivial_1 col_trivial_2 conga_refl) moreover have "D E F CongA D' E' F'" using out546_suma__conga \H Out G I\ assms(4) by blast moreover have "SAMS A B C D E F" using assms(1) by auto ultimately show ?thesis using conga2_sams__sams by auto qed } thus ?thesis using \E Out D F \ \ Bet G H I\ \E Out D F \ SAMS A B C D' E' F'\ by blast next assume "\ Col G H I" show ?thesis proof - have "\ Bet A B C" using Col_def \\ Col A B C\ by auto moreover have "\ J. (C B J CongA D' E' F' \ \ B C OS A J \ \ A B TS C J \ Coplanar A B C J)" proof - have "\ Col A' B' C'" proof - { assume "Col A' B' C'" have "H Out G I \ \ Bet A' B' C'" using SAMS_def assms(5) by blast { assume "H Out G I" then have "False" using Col_cases \\ Col G H I\ out_col by blast } { assume "\ Bet A' B' C'" then have "B' Out A' C'" using not_bet_out \Col A' B' C'\ by blast have "A B C LeA A' B' C'" using assms(1) assms(3) sams_suma__lea123789 by auto then have "B Out A C" using \B' Out A' C'\ out_lea__out by blast then have "Col A B C" using Col_perm out_col by blast then have "False" using \\ Col A B C\ by auto } then have "False" using \H Out G I \ False\ \H Out G I \ \ Bet A' B' C'\ by blast } thus ?thesis by blast qed obtain C1 where P1: "C B C1 CongA D E F \ \ B C OS A C1 \ \ A B TS C C1 \ Coplanar A B C C1" using SAMS_def assms(1) by auto have P1A: "C B C1 CongA D E F" using P1 by simp have P1B: "\ B C OS A C1" using P1 by simp have P1C: "\ A B TS C C1" using P1 by simp have P1D: "Coplanar A B C C1" using P1 by simp have "A B C1 CongA A' B' C'" proof - have "A B C D E F SumA A B C1" using P1A P1B P1D SumA_def \\ Col A B C\ conga_distinct conga_refl not_col_distincts by auto thus ?thesis using assms(3) suma2__conga by auto qed have "SAMS C B C1 G H I" proof - have "D E F CongA C B C1" using P1A not_conga_sym by blast moreover have "G H I CongA G H I" using \\ Col G H I\ conga_refl not_col_distincts by fastforce ultimately show ?thesis using conga2_sams__sams using assms(2) by blast qed then obtain J where T1: "C1 B J CongA G H I \ \ B C1 OS C J \ \ C B TS C1 J \ Coplanar C B C1 J" using SAMS_def by auto have T1A: "C1 B J CongA G H I" using T1 by simp have T1B: "\ B C1 OS C J" using T1 by simp have T1C: "\ C B TS C1 J" using T1 by simp have T1D: "Coplanar C B C1 J" using T1 by simp have "SAMS A B C1 C1 B J" proof - have "A' B' C' CongA A B C1" using \A B C1 CongA A' B' C'\ not_conga_sym by blast moreover have "G H I CongA C1 B J" using T1A not_conga_sym by blast ultimately show ?thesis using assms(5) conga2_sams__sams by auto qed then obtain J' where T2: "C1 B J' CongA C1 B J \ \ B C1 OS A J' \ \ A B TS C1 J' \ Coplanar A B C1 J'" using SAMS_def by auto have T2A: "C1 B J' CongA C1 B J" using T2 by simp have T2B: "\ B C1 OS A J'" using T2 by simp have T2C: "\ A B TS C1 J'" using T2 by simp have T2D: "Coplanar A B C1 J'" using T2 by simp have "A' B' C' CongA A B C1" using \A B C1 CongA A' B' C'\ not_conga_sym by blast then have "\ Col A B C1" using ncol_conga_ncol \\ Col A' B' C'\ by blast have "D E F CongA C B C1" using P1A not_conga_sym by blast then have "\ Col C B C1" using ncol_conga_ncol \\ Col D E F\ by blast then have "Coplanar C1 B A J" using coplanar_trans_1 P1D T1D coplanar_perm_15 coplanar_perm_6 by blast have "Coplanar C1 B J' J" using T2D \Coplanar C1 B A J\ \\ Col A B C1\ coplanar_perm_14 coplanar_perm_6 coplanar_trans_1 by blast have "B Out J' J \ C1 B TS J' J" by (meson T2 \Coplanar C1 B A J\ \\ Col A B C1\ conga_cop__or_out_ts coplanar_trans_1 ncoplanar_perm_14 ncoplanar_perm_6) { assume "B Out J' J" have "\ J. (C B J CongA D' E' F' \ \ B C OS A J \ \ A B TS C J \ Coplanar A B C J)" proof - have "C B C1 C1 B J SumA C B J" proof - have "C1 B J CongA C1 B J" using CongA_def T2A conga_refl by auto moreover have "C B J CongA C B J" using \\ Col C B C1\ calculation conga_diff56 conga_pseudo_refl conga_right_comm not_col_distincts by blast ultimately show ?thesis using T1D T1B SumA_def by blast qed then have "D E F G H I SumA C B J" using conga3_suma__suma by (meson P1A T1A suma2__conga) then have "C B J CongA D' E' F'" using assms(4) suma2__conga by auto moreover have "\ B C OS A J" by (metis (no_types, lifting) Col_perm P1B P1D T1C \\ Col A B C\ \\ Col C B C1\ cop_nos__ts coplanar_perm_8 invert_two_sides l9_2 l9_8_2) moreover have "\ A B TS C J" proof cases assume "Col A B J" thus ?thesis using TS_def invert_two_sides not_col_permutation_3 by blast next assume "\ Col A B J" have "A B OS C J" proof - have "A B OS C C1" by (simp add: P1C P1D \\ Col A B C1\ \\ Col A B C\ cop_nts__os not_col_permutation_2) moreover have "A B OS C1 J" proof - have "A B OS C1 J'" by (metis T2C T2D \B Out J' J\ \\ Col A B C1\ \\ Col A B J\ col_trivial_2 colx cop_nts__os not_col_permutation_2 out_col out_distinct) thus ?thesis using \B Out J' J\ invert_one_side out_out_one_side by blast qed ultimately show ?thesis using one_side_transitivity by blast qed thus ?thesis using l9_9 by blast qed moreover have "Coplanar A B C J" by (meson P1D \Coplanar C1 B A J\ \\ Col A B C1\ coplanar_perm_18 coplanar_perm_2 coplanar_trans_1 not_col_permutation_2) ultimately show ?thesis by blast qed } { assume "C1 B TS J' J" have "B C1 OS C J" proof - have "B C1 TS C J'" proof - have "B C1 TS A J'" by (meson T2B T2D TS_def \C1 B TS J' J\ \\ Col A B C1\ cop_nts__os invert_two_sides ncoplanar_perm_12) moreover have "B C1 OS A C" by (meson P1B P1C P1D \\ Col A B C1\ \\ Col A B C\ \\ Col C B C1\ cop_nts__os invert_one_side ncoplanar_perm_12 not_col_permutation_2 not_col_permutation_3 os_ts1324__os) ultimately show ?thesis using l9_8_2 by blast qed moreover have "B C1 TS J J'" using \C1 B TS J' J\ invert_two_sides l9_2 by blast ultimately show ?thesis using OS_def by blast qed then have "False" by (simp add: T1B) then have "\ J. (C B J CongA D' E' F' \ \ B C OS A J \ \ A B TS C J \ Coplanar A B C J)" by auto } thus ?thesis using \B Out J' J \ \J. C B J CongA D' E' F' \ \ B C OS A J \ \ A B TS C J \ Coplanar A B C J\ \B Out J' J \ C1 B TS J' J\ by blast qed ultimately show ?thesis using SAMS_def not_bet_distincts by auto qed qed qed qed lemma sams_assoc_2: assumes "SAMS A B C D E F" and "SAMS D E F G H I" and "A B C D E F SumA A' B' C'" and "D E F G H I SumA D' E' F'" and "SAMS A B C D' E' F'" shows "SAMS A' B' C' G H I" proof - have "SAMS G H I A' B' C'" proof - have "SAMS G H I D E F" by (simp add: assms(2) sams_sym) moreover have "SAMS D E F A B C" by (simp add: assms(1) sams_sym) moreover have "G H I D E F SumA D' E' F'" by (simp add: assms(4) suma_sym) moreover have "D E F A B C SumA A' B' C'" by (simp add: assms(3) suma_sym) moreover have "SAMS D' E' F' A B C" by (simp add: assms(5) sams_sym) ultimately show ?thesis using sams_assoc_1 by blast qed thus ?thesis using sams_sym by blast qed lemma sams_assoc: assumes "SAMS A B C D E F" and "SAMS D E F G H I" and "A B C D E F SumA A' B' C'" and "D E F G H I SumA D' E' F'" shows "(SAMS A' B' C' G H I) \ (SAMS A B C D' E' F')" using sams_assoc_1 sams_assoc_2 by (meson assms(1) assms(2) assms(3) assms(4)) lemma sams_nos__nts: assumes "SAMS A B C C B J" and "\ B C OS A J" shows "\ A B TS C J" proof - obtain J' where P1: "C B J' CongA C B J \ \ B C OS A J' \ \ A B TS C J' \ Coplanar A B C J'" using SAMS_def assms(1) by blast have P1A: "C B J' CongA C B J" using P1 by simp have P1B: "\ B C OS A J'" using P1 by simp have P1C: "\ A B TS C J'" using P1 by simp have P1D: "Coplanar A B C J'" using P1 by simp have P2: "B Out C J \ \ Bet A B C" using SAMS_def assms(1) by blast { assume "A B TS C J" have "Coplanar C B J J'" proof - have "\ Col A C B" using TS_def \A B TS C J\ not_col_permutation_4 by blast moreover have "Coplanar A C B J" using \A B TS C J\ ncoplanar_perm_2 ts__coplanar by blast moreover have "Coplanar A C B J'" using P1D ncoplanar_perm_2 by blast ultimately show ?thesis using coplanar_trans_1 by blast qed have "B Out J J' \ C B TS J J'" by (metis P1 \A B TS C J\ \Coplanar C B J J'\ bet_conga__bet bet_out col_conga_col col_two_sides_bet conga_distinct conga_os__out conga_sym cop_nts__os invert_two_sides l5_2 l6_6 not_col_permutation_3 not_col_permutation_4) { assume "B Out J J'" have "\ Col B A J \ \ Col B A J'" using TS_def \A B TS C J\ not_col_permutation_3 by blast then have "A B OS C J'" by (metis (full_types) \B Out J J'\ Col_cases P1C P1D TS_def \A B TS C J\ col2__eq cop_nts__os l6_3_1 out_col) then have "A B TS C J'" by (meson \A B TS C J\ \B Out J J'\ l6_6 l9_2 not_col_distincts out_two_sides_two_sides) then have "False" by (simp add: P1C) } { assume "C B TS J J'" then have "\ Col C A B \ \ Col J A B" using TS_def \A B TS C J\ by blast then have "False" by (metis P1B P1D TS_def \C B TS J J'\ assms(2) cop_nts__os invert_two_sides l9_8_1 ncoplanar_perm_12 not_col_permutation_1) } then have "False" using \B Out J J' \ False\ \B Out J J' \ C B TS J J'\ by blast } thus ?thesis by auto qed lemma conga_sams_nos__nts: assumes "SAMS A B C D E F" and "C B J CongA D E F" and "\ B C OS A J" shows "\ A B TS C J" proof - have "SAMS A B C C B J" proof - have "A B C CongA A B C" using assms(1) conga_refl sams_distincts by fastforce moreover have "D E F CongA C B J" using assms(2) not_conga_sym by blast ultimately show ?thesis using assms(1) conga2_sams__sams by auto qed thus ?thesis by (simp add: assms(3) sams_nos__nts) qed lemma sams_lea2_suma2__conga123: assumes "A B C LeA A' B' C'" and "D E F LeA D' E' F'" and "SAMS A' B' C' D' E' F'" and "A B C D E F SumA G H I" and "A' B' C' D' E' F' SumA G H I" shows "A B C CongA A' B' C'" proof - have "SAMS A B C D E F" using assms(1) assms(2) assms(3) sams_lea2__sams by blast moreover have "SAMS A' B' C' D E F" by (metis assms(2) assms(3) lea_refl sams_distincts sams_lea2__sams) moreover have "A' B' C' D E F SumA G H I" proof - obtain G' H' I' where P1: "A' B' C' D E F SumA G' H' I'" using calculation(2) ex_suma sams_distincts by blast show ?thesis proof - have "A' \ B' \ B' \ C'" using assms(1) lea_distincts by blast then have "A' B' C' CongA A' B' C'" using conga_refl by auto moreover have "D \ E \ E \ F" using \SAMS A B C D E F\ sams_distincts by blast then have "D E F CongA D E F" using conga_refl by auto moreover have "G' H' I' CongA G H I" proof - have "G' H' I' LeA G H I" using P1 assms(2) assms(3) assms(5) sams_lea456_suma2__lea by blast moreover have "G H I LeA G' H' I'" proof - have "SAMS A' B' C' D E F" using \SAMS A' B' C' D E F\ by auto thus ?thesis using P1 assms(1) assms(4) sams_lea123_suma2__lea by blast qed ultimately show ?thesis by (simp add: lea_asym) qed ultimately show ?thesis using P1 conga3_suma__suma by blast qed qed ultimately show ?thesis using assms(4) sams2_suma2__conga123 by blast qed lemma sams_lea2_suma2__conga456: assumes "A B C LeA A' B' C'" and "D E F LeA D' E' F'" and "SAMS A' B' C' D' E' F'" and "A B C D E F SumA G H I" and "A' B' C' D' E' F' SumA G H I" shows "D E F CongA D' E' F'" proof - have "SAMS D' E' F' A' B' C'" by (simp add: assms(3) sams_sym) moreover have "D E F A B C SumA G H I" by (simp add: assms(4) suma_sym) moreover have "D' E' F' A' B' C' SumA G H I" by (simp add: assms(5) suma_sym) ultimately show ?thesis using assms(1) assms(2) sams_lea2_suma2__conga123 by auto qed lemma sams_suma__out213: assumes "A B C D E F SumA D E F" and "SAMS A B C D E F" shows "B Out A C" proof - have "E \ D" using assms(2) sams_distincts by blast then have "E Out D D" using out_trivial by auto moreover have "D E D CongA A B C" proof - have "D E D LeA A B C" using assms(1) lea121345 suma_distincts by auto moreover have "E \ D \ E \ F" using assms(2) sams_distincts by blast then have "D E F LeA D E F" using lea_refl by auto moreover have "D E D D E F SumA D E F" proof - have "\ E D OS D F" using os_distincts by auto moreover have "Coplanar D E D F" using ncop_distincts by auto ultimately show ?thesis using SumA_def \D E F LeA D E F\ lea_asym by blast qed ultimately show ?thesis using assms(1) assms(2) sams_lea2_suma2__conga123 by auto qed ultimately show ?thesis using eq_conga_out by blast qed lemma sams_suma__out546: assumes "A B C D E F SumA A B C" and "SAMS A B C D E F" shows "E Out D F" proof - have "D E F A B C SumA A B C" using assms(1) suma_sym by blast moreover have "SAMS D E F A B C" using assms(2) sams_sym by blast ultimately show ?thesis using sams_suma__out213 by blast qed lemma sams_lea_lta123_suma2__lta: assumes "A B C LtA A' B' C'" and "D E F LeA D' E' F'" and "SAMS A' B' C' D' E' F'" and "A B C D E F SumA G H I" and "A' B' C' D' E' F' SumA G' H' I'" shows "G H I LtA G' H' I'" proof - have "G H I LeA G' H' I'" proof - have "A B C LeA A' B' C'" by (simp add: assms(1) lta__lea) thus ?thesis using assms(2) assms(3) assms(4) assms(5) sams_lea2_suma2__lea by blast qed moreover have "\ G H I CongA G' H' I'" proof - { assume "G H I CongA G' H' I'" have "A B C CongA A' B' C'" proof - have "A B C LeA A' B' C'" by (simp add: assms(1) lta__lea) moreover have "A' B' C' D' E' F' SumA G H I" by (meson \G H I CongA G' H' I'\ assms(3) assms(5) conga3_suma__suma conga_sym sams2_suma2__conga123 sams2_suma2__conga456) ultimately show ?thesis using assms(2) assms(3) assms(4) sams_lea2_suma2__conga123 by blast qed then have "False" using assms(1) lta_not_conga by auto } thus ?thesis by auto qed ultimately show ?thesis using LtA_def by blast qed lemma sams_lea_lta456_suma2__lta: assumes "A B C LeA A' B' C'" and "D E F LtA D' E' F'" and "SAMS A' B' C' D' E' F'" and "A B C D E F SumA G H I" and "A' B' C' D' E' F' SumA G' H' I'" shows "G H I LtA G' H' I'" using sams_lea_lta123_suma2__lta by (meson assms(1) assms(2) assms(3) assms(4) assms(5) sams_sym suma_sym) lemma sams_lta2_suma2__lta: assumes "A B C LtA A' B' C'" and "D E F LtA D' E' F'" and "SAMS A' B' C' D' E' F'" and "A B C D E F SumA G H I" and "A' B' C' D' E' F' SumA G' H' I'" shows "G H I LtA G' H' I'" using sams_lea_lta123_suma2__lta by (meson LtA_def assms(1) assms(2) assms(3) assms(4) assms(5)) lemma sams_lea2_suma2__lea123: assumes "D' E' F' LeA D E F" and "G H I LeA G' H' I'" and "SAMS A B C D E F" and "A B C D E F SumA G H I" and "A' B' C' D' E' F' SumA G' H' I'" shows "A B C LeA A' B' C'" proof cases assume "A' B' C' LtA A B C" then have "G' H' I' LtA G H I" using sams_lea_lta123_suma2__lta using assms(1) assms(3) assms(4) assms(5) by blast then have "\ G H I LeA G' H' I'" using lea__nlta by blast then have "False" using assms(2) by auto thus ?thesis by auto next assume "\ A' B' C' LtA A B C" have "A' \ B' \ B' \ C' \ A \ B \ B \ C" using assms(4) assms(5) suma_distincts by auto thus ?thesis by (simp add: \\ A' B' C' LtA A B C\ nlta__lea) qed lemma sams_lea2_suma2__lea456: assumes "A' B' C' LeA A B C" and "G H I LeA G' H' I'" and "SAMS A B C D E F" and "A B C D E F SumA G H I" and "A' B' C' D' E' F' SumA G' H' I'" shows "D E F LeA D' E' F'" proof - have "SAMS D E F A B C" by (simp add: assms(3) sams_sym) moreover have "D E F A B C SumA G H I" by (simp add: assms(4) suma_sym) moreover have "D' E' F' A' B' C' SumA G' H' I'" by (simp add: assms(5) suma_sym) ultimately show ?thesis using assms(1) assms(2) sams_lea2_suma2__lea123 by blast qed lemma sams_lea_lta456_suma2__lta123: assumes "D' E' F' LtA D E F" and "G H I LeA G' H' I'" and "SAMS A B C D E F" and "A B C D E F SumA G H I" and "A' B' C' D' E' F' SumA G' H' I'" shows "A B C LtA A' B' C'" proof cases assume "A' B' C' LeA A B C" then have "G' H' I' LtA G H I" using sams_lea_lta456_suma2__lta assms(1) assms(3) assms(4) assms(5) by blast then have "\ G H I LeA G' H' I'" using lea__nlta by blast then have "False" using assms(2) by blast thus ?thesis by blast next assume "\ A' B' C' LeA A B C" have "A' \ B' \ B' \ C' \ A \ B \ B \ C" using assms(4) assms(5) suma_distincts by auto thus ?thesis using nlea__lta by (simp add: \\ A' B' C' LeA A B C\) qed lemma sams_lea_lta123_suma2__lta456: assumes "A' B' C' LtA A B C" and "G H I LeA G' H' I'" and "SAMS A B C D E F" and "A B C D E F SumA G H I" and "A' B' C' D' E' F' SumA G' H' I'" shows "D E F LtA D' E' F'" proof - have "SAMS D E F A B C" by (simp add: assms(3) sams_sym) moreover have "D E F A B C SumA G H I" by (simp add: assms(4) suma_sym) moreover have "D' E' F' A' B' C' SumA G' H' I'" by (simp add: assms(5) suma_sym) ultimately show ?thesis using assms(1) assms(2) sams_lea_lta456_suma2__lta123 by blast qed lemma sams_lea_lta789_suma2__lta123: assumes "D' E' F' LeA D E F" and "G H I LtA G' H' I'" and "SAMS A B C D E F" and "A B C D E F SumA G H I" and "A' B' C' D' E' F' SumA G' H' I'" shows "A B C LtA A' B' C'" proof cases assume "A' B' C' LeA A B C" then have "G' H' I' LeA G H I" using assms(1) assms(3) assms(4) assms(5) sams_lea2_suma2__lea by blast then have "\ G H I LtA G' H' I'" by (simp add: lea__nlta) then have "False" using assms(2) by blast thus ?thesis by auto next assume "\ A' B' C' LeA A B C" have "A' \ B' \ B' \ C' \ A \ B \ B \ C" using assms(4) assms(5) suma_distincts by auto thus ?thesis using nlea__lta by (simp add: \\ A' B' C' LeA A B C\) qed lemma sams_lea_lta789_suma2__lta456: assumes "A' B' C' LeA A B C" and "G H I LtA G' H' I'" and "SAMS A B C D E F" and "A B C D E F SumA G H I" and "A' B' C' D' E' F' SumA G' H' I'" shows "D E F LtA D' E' F'" proof - have "SAMS D E F A B C" by (simp add: assms(3) sams_sym) moreover have "D E F A B C SumA G H I" by (simp add: assms(4) suma_sym) moreover have "D' E' F' A' B' C' SumA G' H' I'" using assms(5) suma_sym by blast ultimately show ?thesis using assms(1) assms(2) sams_lea_lta789_suma2__lta123 by blast qed lemma sams_lta2_suma2__lta123: assumes "D' E' F' LtA D E F" and "G H I LtA G' H' I'" and "SAMS A B C D E F" and "A B C D E F SumA G H I" and "A' B' C' D' E' F' SumA G' H' I'" shows "A B C LtA A' B' C'" proof - have "D' E' F' LeA D E F" by (simp add: assms(1) lta__lea) thus ?thesis using assms(2) assms(3) assms(4) assms(5) sams_lea_lta789_suma2__lta123 by blast qed lemma sams_lta2_suma2__lta456: assumes "A' B' C' LtA A B C" and "G H I LtA G' H' I'" and "SAMS A B C D E F" and "A B C D E F SumA G H I" and "A' B' C' D' E' F' SumA G' H' I'" shows "D E F LtA D' E' F'" proof - have "A' B' C' LeA A B C" by (simp add: assms(1) lta__lea) thus ?thesis using assms(2) assms(3) assms(4) assms(5) sams_lea_lta789_suma2__lta456 by blast qed lemma sams123231: assumes "A \ B" and "A \ C" and "B \ C" shows "SAMS A B C B C A" proof - obtain A' where "B Midpoint A A'" using symmetric_point_construction by auto then have "A' \ B" using assms(1) midpoint_not_midpoint by blast moreover have "Bet A B A'" by (simp add: \B Midpoint A A'\ midpoint_bet) moreover have "B C A LeA C B A'" proof cases assume "Col A B C" show ?thesis proof cases assume "Bet A C B" thus ?thesis by (metis assms(2) assms(3) between_exchange3 calculation(1) calculation(2) l11_31_2) next assume "\ Bet A C B" then have "C Out B A" using Col_cases \Col A B C\ l6_6 or_bet_out by blast thus ?thesis using assms(3) calculation(1) l11_31_1 by auto qed next assume "\ Col A B C" thus ?thesis using l11_41_aux \B Midpoint A A'\ calculation(1) lta__lea midpoint_bet not_col_permutation_4 by blast qed ultimately show ?thesis using assms(1) sams_chara by blast qed lemma col_suma__col: assumes "Col D E F" and "A B C B C A SumA D E F" shows "Col A B C" proof - { assume "\ Col A B C" have "False" proof cases assume "Bet D E F" obtain P where P1: "Bet A B P \ Cong A B B P" using Cong_perm segment_construction by blast have "D E F LtA A B P" proof - have "A B C LeA A B C" using \\ Col A B C\ lea_total not_col_distincts by blast moreover have "B C TS A P" by (metis Cong_perm P1 \\ Col A B C\ bet__ts bet_col between_trivial2 cong_reverse_identity not_col_permutation_1) then have "B C A LtA C B P" using Col_perm P1 \\ Col A B C\ calculation l11_41_aux ts_distincts by blast moreover have "A B C C B P SumA A B P" by (simp add: \B C TS A P\ ts__suma_1) ultimately show ?thesis by (meson P1 Tarski_neutral_dimensionless.sams_lea_lta456_suma2__lta Tarski_neutral_dimensionless_axioms assms(2) bet_suma__sams) qed thus ?thesis using P1 \Bet D E F\ bet2_lta__lta lta_distincts by blast next assume "\ Bet D E F" have "C Out B A" proof - have "E Out D F" by (simp add: \\ Bet D E F\ assms(1) l6_4_2) moreover have "B C A LeA D E F" using sams_suma__lea456789 by (metis assms(2) sams123231 suma_distincts) ultimately show ?thesis using out_lea__out by blast qed thus ?thesis using Col_cases \\ Col A B C\ out_col by blast qed } thus ?thesis by auto qed lemma ncol_suma__ncol: assumes "\ Col A B C" and "A B C B C A SumA D E F" shows "\ Col D E F" using col_suma__col assms(1) assms(2) by blast lemma per2_suma__bet: assumes "Per A B C" and "Per D E F" and "A B C D E F SumA G H I" shows "Bet G H I" proof - obtain A1 where P1: "C B A1 CongA D E F \ \ B C OS A A1 \ Coplanar A B C A1 \ A B A1 CongA G H I" using SumA_def assms(3) by blast then have "D E F CongA A1 B C" using conga_right_comm conga_sym by blast then have "Per A1 B C" using assms(2) l11_17 by blast have "Bet A B A1" proof - have "Col B A A1" proof - have "Coplanar C A A1 B" using P1 ncoplanar_perm_10 by blast moreover have "C \ B" using \D E F CongA A1 B C\ conga_distinct by auto ultimately show ?thesis using assms(1) \Per A1 B C\ col_permutation_2 cop_per2__col by blast qed moreover have "B C TS A A1" proof - have "Coplanar B C A A1" using calculation ncop__ncols by blast moreover have "A \ B \ B \ C" using CongA_def P1 by blast then have "\ Col A B C" by (simp add: assms(1) per_not_col) moreover have "A1 \ B \ B \ C" using \D E F CongA A1 B C\ conga_distinct by blast then have "\ Col A1 B C" using \Per A1 B C\ by (simp add: per_not_col) ultimately show ?thesis by (simp add: P1 cop_nos__ts) qed ultimately show ?thesis using col_two_sides_bet by blast qed thus ?thesis using P1 bet_conga__bet by blast qed lemma bet_per2__suma: assumes "A \ B" and "B \ C" and "D \ E" and "E \ F" and "G \ H" and "H \ I" and "Per A B C" and "Per D E F" and "Bet G H I" shows "A B C D E F SumA G H I" proof - obtain G' H' I' where "A B C D E F SumA G' H' I'" using assms(1) assms(2) assms(3) assms(4) ex_suma by blast moreover have "A B C CongA A B C" using assms(1) assms(2) conga_refl by auto moreover have "D E F CongA D E F" using assms(3) assms(4) conga_refl by auto moreover have "G' H' I' CongA G H I" proof - have "G' \ H'" using calculation(1) suma_distincts by auto moreover have "H' \ I'" using \A B C D E F SumA G' H' I'\ suma_distincts by blast moreover have "Bet G' H' I'" using \A B C D E F SumA G' H' I'\ assms(7) assms(8) per2_suma__bet by auto ultimately show ?thesis using conga_line by (simp add: assms(5) assms(6) assms(9)) qed ultimately show ?thesis using conga3_suma__suma by blast qed lemma per2__sams: assumes "A \ B" and "B \ C" and "D \ E" and "E \ F" and "Per A B C" and "Per D E F" shows "SAMS A B C D E F" proof - obtain G H I where "A B C D E F SumA G H I" using assms(1) assms(2) assms(3) assms(4) ex_suma by blast moreover then have "Bet G H I" using assms(5) assms(6) per2_suma__bet by auto ultimately show ?thesis using bet_suma__sams by blast qed lemma bet_per_suma__per456: assumes "Per A B C" and "Bet G H I" and "A B C D E F SumA G H I" shows "Per D E F" proof - obtain A1 where "B Midpoint A A1" using symmetric_point_construction by auto have "\ Col A B C" using assms(1) assms(3) per_col_eq suma_distincts by blast have "A B C CongA D E F" proof - have "SAMS A B C A B C" using \\ Col A B C\ assms(1) not_col_distincts per2__sams by auto moreover have "SAMS A B C D E F" using assms(2) assms(3) bet_suma__sams by blast moreover have "A B C A B C SumA G H I" using assms(1) assms(2) assms(3) bet_per2__suma suma_distincts by blast ultimately show ?thesis using assms(3) sams2_suma2__conga456 by auto qed thus ?thesis using assms(1) l11_17 by blast qed lemma bet_per_suma__per123: assumes "Per D E F" and "Bet G H I" and "A B C D E F SumA G H I" shows "Per A B C" using bet_per_suma__per456 by (meson assms(1) assms(2) assms(3) suma_sym) lemma bet_suma__per: assumes "Bet D E F" and "A B C A B C SumA D E F" shows "Per A B C" proof - obtain A' where "C B A' CongA A B C \ A B A' CongA D E F" using SumA_def assms(2) by blast have "Per C B A" proof - have "Bet A B A'" proof - have "D E F CongA A B A'" using \C B A' CongA A B C \ A B A' CongA D E F\ not_conga_sym by blast thus ?thesis using assms(1) bet_conga__bet by blast qed moreover have "C B A CongA C B A'" using conga_left_comm not_conga_sym \C B A' CongA A B C \ A B A' CongA D E F\ by blast ultimately show ?thesis using l11_18_2 by auto qed thus ?thesis using Per_cases by auto qed lemma acute__sams: assumes "Acute A B C" shows "SAMS A B C A B C" proof - obtain A' where "B Midpoint A A'" using symmetric_point_construction by auto show ?thesis proof - have "A \ B \ A' \ B" using \B Midpoint A A'\ acute_distincts assms is_midpoint_id_2 by blast moreover have "Bet A B A'" by (simp add: \B Midpoint A A'\ midpoint_bet) moreover have "Obtuse C B A'" using acute_bet__obtuse assms calculation(1) calculation(2) obtuse_sym by blast then have "A B C LeA C B A'" by (metis acute__not_obtuse assms calculation(1) lea_obtuse_obtuse lea_total obtuse_distincts) ultimately show ?thesis using sams_chara by blast qed qed lemma acute_suma__nbet: assumes "Acute A B C" and "A B C A B C SumA D E F" shows "\ Bet D E F" proof - { assume "Bet D E F" then have "Per A B C" using assms(2) bet_suma__per by auto then have "A B C LtA A B C" using acute_not_per assms(1) by auto then have "False" by (simp add: nlta) } thus ?thesis by blast qed lemma acute2__sams: assumes "Acute A B C" and "Acute D E F" shows "SAMS A B C D E F" by (metis acute__sams acute_distincts assms(1) assms(2) lea_total sams_lea2__sams) lemma acute2_suma__nbet_a: assumes "Acute A B C" and "D E F LeA A B C" and "A B C D E F SumA G H I" shows "\ Bet G H I" proof - { assume "Bet G H I" obtain A' B' C' where "A B C A B C SumA A' B' C'" using acute_distincts assms(1) ex_suma by moura have "G H I LeA A' B' C'" proof - have "A B C LeA A B C" using acute_distincts assms(1) lea_refl by blast moreover have "SAMS A B C A B C" by (simp add: acute__sams assms(1)) ultimately show ?thesis using \A B C A B C SumA A' B' C'\ assms(1) assms(2) assms(3) sams_lea456_suma2__lea by blast qed then have "Bet A' B' C'" using \Bet G H I\ bet_lea__bet by blast then have "False" using acute_suma__nbet assms(1) assms(3) \A B C A B C SumA A' B' C'\ by blast } thus ?thesis by blast qed lemma acute2_suma__nbet: assumes "Acute A B C" and "Acute D E F" and "A B C D E F SumA G H I" shows "\ Bet G H I" proof - have "A \ B \ B \ C \ D \ E \ E \ F" using assms(3) suma_distincts by auto then have "A B C LeA D E F \ D E F LeA A B C" by (simp add: lea_total) moreover { assume P3: "A B C LeA D E F" have "D E F A B C SumA G H I" by (simp add: assms(3) suma_sym) then have "\ Bet G H I" using P3 assms(2) acute2_suma__nbet_a by auto } { assume "D E F LeA A B C" then have "\ Bet G H I" using acute2_suma__nbet_a assms(1) assms(3) by auto } thus ?thesis using \A B C LeA D E F \ \ Bet G H I\ calculation by blast qed lemma acute_per__sams: assumes "A \ B" and "B \ C" and "Per A B C" and "Acute D E F" shows "SAMS A B C D E F" proof - have "SAMS A B C A B C" by (simp add: assms(1) assms(2) assms(3) per2__sams) moreover have "A B C LeA A B C" using assms(1) assms(2) lea_refl by auto moreover have "D E F LeA A B C" by (metis acute_distincts acute_lea_acute acute_not_per assms(1) assms(2) assms(3) assms(4) lea_total) ultimately show ?thesis using sams_lea2__sams by blast qed lemma acute_per_suma__nbet: assumes "A \ B" and "B \ C" and "Per A B C" and "Acute D E F" and "A B C D E F SumA G H I" shows "\ Bet G H I" proof - { assume "Bet G H I" have "G H I LtA G H I" proof - have "A B C LeA A B C" using assms(1) assms(2) lea_refl by auto moreover have "D E F LtA A B C" by (simp add: acute_per__lta assms(1) assms(2) assms(3) assms(4)) moreover have "SAMS A B C A B C" by (simp add: assms(1) assms(2) assms(3) per2__sams) moreover have "A B C D E F SumA G H I" by (simp add: assms(5)) moreover have "A B C A B C SumA G H I" by (meson Tarski_neutral_dimensionless.bet_per_suma__per456 Tarski_neutral_dimensionless_axioms \Bet G H I\ acute_not_per assms(3) assms(4) calculation(4)) ultimately show ?thesis using sams_lea_lta456_suma2__lta by blast qed then have "False" by (simp add: nlta) } thus ?thesis by blast qed lemma obtuse__nsams: assumes "Obtuse A B C" shows "\ SAMS A B C A B C" proof - { assume "SAMS A B C A B C" obtain A' where "B Midpoint A A'" using symmetric_point_construction by auto have "A B C LtA A B C" proof - have "A B C LeA A' B C" by (metis \B Midpoint A A'\ \SAMS A B C A B C\ lea_right_comm midpoint_bet midpoint_distinct_2 sams_chara sams_distincts) moreover have "A' B C LtA A B C" using \B Midpoint A A'\ assms calculation lea_distincts midpoint_bet obtuse_chara by blast ultimately show ?thesis using lea__nlta by blast qed then have "False" by (simp add: nlta) } thus ?thesis by blast qed lemma nbet_sams_suma__acute: assumes "\ Bet D E F" and "SAMS A B C A B C" and "A B C A B C SumA D E F" shows "Acute A B C" proof - have "Acute A B C \ Per A B C \ Obtuse A B C" by (metis angle_partition l8_20_1_R1 l8_5) { assume "Per A B C" then have "Bet D E F" using assms(3) per2_suma__bet by auto then have "False" using assms(1) by auto } { assume "Obtuse A B C" then have "\ SAMS A B C A B C" by (simp add: obtuse__nsams) then have "False" using assms(2) by auto } thus ?thesis using \Acute A B C \ Per A B C \ Obtuse A B C\ \Per A B C \ False\ by auto qed lemma nsams__obtuse: assumes "A \ B" and "B \ C" and "\ SAMS A B C A B C" shows "Obtuse A B C" proof - { assume "Per A B C" obtain A' where "B Midpoint A A'" using symmetric_point_construction by blast have "\ Col A B C" using \Per A B C\ assms(1) assms(2) per_col_eq by blast have "SAMS A B C A B C" proof - have "C B A' CongA A B C" using \Per A B C\ assms(1) assms(2) assms(3) per2__sams by blast moreover have "\ B C OS A A'" by (meson Col_cases \B Midpoint A A'\ col_one_side_out l6_4_1 midpoint_bet midpoint_col) moreover have "\ A B TS C A'" using Col_def Midpoint_def TS_def \B Midpoint A A'\ by blast moreover have "Coplanar A B C A'" using \Per A B C\ assms(1) assms(2) assms(3) per2__sams by blast ultimately show ?thesis using SAMS_def \\ Col A B C\ assms(1) bet_col by auto qed then have "False" using assms(3) by auto } { assume "Acute A B C" then have "SAMS A B C A B C" by (simp add: acute__sams) then have "False" using assms(3) by auto } thus ?thesis using \Per A B C \ False\ angle_partition assms(1) assms(2) by auto qed lemma sams2_suma2__conga: assumes "SAMS A B C A B C" and "A B C A B C SumA D E F" and "SAMS A' B' C' A' B' C'" and "A' B' C' A' B' C' SumA D E F" shows "A B C CongA A' B' C'" proof - have "A B C LeA A' B' C' \ A' B' C' LeA A B C" using assms(1) assms(3) lea_total sams_distincts by auto moreover have "A B C LeA A' B' C' \ A B C CongA A' B' C'" using assms(2) assms(3) assms(4) sams_lea2_suma2__conga123 by auto ultimately show ?thesis by (meson Tarski_neutral_dimensionless.conga_sym Tarski_neutral_dimensionless.sams_lea2_suma2__conga123 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(4)) qed lemma acute2_suma2__conga: assumes "Acute A B C" and "A B C A B C SumA D E F" and "Acute A' B' C'" and "A' B' C' A' B' C' SumA D E F" shows "A B C CongA A' B' C'" proof - have "SAMS A B C A B C" by (simp add: acute__sams assms(1)) moreover have "SAMS A' B' C' A' B' C'" by (simp add: acute__sams assms(3)) ultimately show ?thesis using assms(2) assms(4) sams2_suma2__conga by auto qed lemma bet2_suma__out: assumes "Bet A B C" and "Bet D E F" and "A B C D E F SumA G H I" shows "H Out G I" proof - have "A B C D E F SumA A B A" proof - have "C B A CongA D E F" by (metis Bet_cases assms(1) assms(2) assms(3) conga_line suma_distincts) moreover have "\ B C OS A A" by (simp add: Col_def assms(1) col124__nos) moreover have "Coplanar A B C A" using ncop_distincts by blast moreover have "A B A CongA A B A" using calculation(1) conga_diff2 conga_trivial_1 by auto ultimately show ?thesis using SumA_def by blast qed then have "A B A CongA G H I" using assms(3) suma2__conga by auto thus ?thesis using eq_conga_out by auto qed lemma col2_suma__col: assumes "Col A B C" and "Col D E F" and "A B C D E F SumA G H I" shows "Col G H I" proof cases assume "Bet A B C" show ?thesis proof cases assume "Bet D E F" thus ?thesis using bet2_suma__out by (meson \Bet A B C\ assms(3) not_col_permutation_4 out_col) next assume "\ Bet D E F" show ?thesis proof - have "E Out D F" using \\ Bet D E F\ assms(2) or_bet_out by auto then have "A B C CongA G H I" using assms(3) out546_suma__conga by auto thus ?thesis using assms(1) col_conga_col by blast qed qed next assume "\ Bet A B C" have "D E F CongA G H I" proof - have "B Out A C" by (simp add: \\ Bet A B C\ assms(1) l6_4_2) thus ?thesis using assms(3) out213_suma__conga by auto qed thus ?thesis using assms(2) col_conga_col by blast qed lemma suma_suppa__bet: assumes "A B C SuppA D E F" and "A B C D E F SumA G H I" shows "Bet G H I" proof - obtain A' where P1: "Bet A B A' \ D E F CongA C B A'" using SuppA_def assms(1) by auto obtain A'' where P2: "C B A'' CongA D E F \ \ B C OS A A'' \ Coplanar A B C A'' \ A B A'' CongA G H I" using SumA_def assms(2) by auto have "B Out A' A'' \ C B TS A' A''" proof - have "Coplanar C B A' A''" proof - have "Coplanar C A'' B A" using P2 coplanar_perm_17 by blast moreover have "B \ A" using SuppA_def assms(1) by auto moreover have "Col B A A'" using P1 by (simp add: bet_col col_permutation_4) ultimately show ?thesis using col_cop__cop coplanar_perm_3 by blast qed moreover have "C B A' CongA C B A''" proof - have "C B A' CongA D E F" using P1 not_conga_sym by blast moreover have "D E F CongA C B A''" using P2 not_conga_sym by blast ultimately show ?thesis using not_conga by blast qed ultimately show ?thesis using conga_cop__or_out_ts by simp qed have "Bet A B A''" proof - have "\ C B TS A' A''" proof - { assume "C B TS A' A''" have "B C TS A A'" proof - { assume "Col A B C" then have "Col A' C B" using P1 assms(1) bet_col bet_col1 col3 suppa_distincts by blast then have "False" using TS_def \C B TS A' A''\ by auto } then have "\ Col A B C" by auto moreover have "\ Col A' B C" using TS_def \C B TS A' A''\ not_col_permutation_5 by blast moreover have "\ T. (Col T B C \ Bet A T A')" using P1 not_col_distincts by blast ultimately show ?thesis by (simp add: TS_def) qed then have "B C OS A A''" using OS_def \C B TS A' A''\ invert_two_sides l9_2 by blast then have "False" using P2 by simp } thus ?thesis by blast qed then have "B Out A' A''" using \B Out A' A'' \ C B TS A' A''\ by auto moreover have "A' \ B \ A'' \ B \ A \ B" using P2 calculation conga_diff1 out_diff1 out_diff2 by blast moreover have "Bet A' B A" using P1 Bet_perm by blast ultimately show ?thesis using bet_out__bet between_symmetry by blast qed moreover have "A B A'' CongA G H I" using P2 by blast ultimately show ?thesis using bet_conga__bet by blast qed lemma bet_suppa__suma: assumes "G \ H" and "H \ I" and "A B C SuppA D E F" and "Bet G H I" shows "A B C D E F SumA G H I" proof - obtain G' H' I' where "A B C D E F SumA G' H' I'" using assms(3) ex_suma suppa_distincts by blast moreover have "A B C CongA A B C" using calculation conga_refl suma_distincts by fastforce moreover have "D \ E \ E \ F" using assms(3) suppa_distincts by auto then have "D E F CongA D E F" using conga_refl by auto moreover have "G' H' I' CongA G H I" proof - have "G' \ H' \ H' \ I'" using calculation(1) suma_distincts by auto moreover have "Bet G' H' I'" using \A B C D E F SumA G' H' I'\ assms(3) suma_suppa__bet by blast ultimately show ?thesis using assms(1) assms(2) assms(4) conga_line by auto qed ultimately show ?thesis using conga3_suma__suma by blast qed lemma bet_suma__suppa: assumes "A B C D E F SumA G H I" and "Bet G H I" shows "A B C SuppA D E F" proof - obtain A' where "C B A' CongA D E F \ A B A' CongA G H I" using SumA_def assms(1) by blast moreover have "G H I CongA A B A'" using calculation not_conga_sym by blast then have "Bet A B A'" using assms(2) bet_conga__bet by blast moreover have "D E F CongA C B A'" using calculation(1) not_conga_sym by blast ultimately show ?thesis by (metis SuppA_def conga_diff1) qed lemma bet2_suma__suma: assumes "A' \ B" and "D' \ E" and "Bet A B A'" and "Bet D E D'" and "A B C D E F SumA G H I" shows "A' B C D' E F SumA G H I" proof - obtain J where P1: "C B J CongA D E F \ \ B C OS A J \ Coplanar A B C J \ A B J CongA G H I" using SumA_def assms(5) by auto moreover obtain C' where P2: "Bet C B C' \ Cong B C' B C" using segment_construction by blast moreover have "A B C' D' E F SumA G H I" proof - have "C' B J CongA D' E F" by (metis assms(2) assms(4) calculation(1) calculation(2) cong_diff_3 conga_diff1 l11_13) moreover have "\ B C' OS A J" by (metis Col_cases P1 P2 bet_col col_one_side cong_diff) moreover have "Coplanar A B C' J" by (smt P1 P2 bet_col bet_col1 col2_cop__cop cong_diff ncoplanar_perm_5) ultimately show ?thesis using P1 SumA_def by blast qed moreover have "A B C' CongA A' B C" using assms(1) assms(3) assms(5) between_symmetry calculation(2) calculation(3) l11_14 suma_distincts by auto moreover have "D' \ E \ E \ F" using assms(2) calculation(1) conga_distinct by blast then have "D' E F CongA D' E F" using conga_refl by auto moreover have "G \ H \ H \ I" using assms(5) suma_distincts by blast then have "G H I CongA G H I" using conga_refl by auto ultimately show ?thesis using conga3_suma__suma by blast qed lemma suma_suppa2__suma: assumes "A B C SuppA A' B' C'" and "D E F SuppA D' E' F'" and "A B C D E F SumA G H I" shows "A' B' C' D' E' F' SumA G H I" proof - obtain A0 where P1: "Bet A B A0 \ A' B' C' CongA C B A0" using SuppA_def assms(1) by auto obtain D0 where P2: "Bet D E D0 \ D' E' F' CongA F E D0" using SuppA_def assms(2) by auto show ?thesis proof - have "A0 B C D0 E F SumA G H I" proof - have "A0 \ B" using CongA_def P1 by auto moreover have "D0 \ E" using CongA_def P2 by blast ultimately show ?thesis using P1 P2 assms(3) bet2_suma__suma by auto qed moreover have "A0 B C CongA A' B' C'" using P1 conga_left_comm not_conga_sym by blast moreover have "D0 E F CongA D' E' F'" using P2 conga_left_comm not_conga_sym by blast moreover have "G \ H \ H \ I" using assms(3) suma_distincts by blast then have "G H I CongA G H I" using conga_refl by auto ultimately show ?thesis using conga3_suma__suma by blast qed qed lemma suma2_obtuse2__conga: assumes "Obtuse A B C" and "A B C A B C SumA D E F" and "Obtuse A' B' C'" and "A' B' C' A' B' C' SumA D E F" shows "A B C CongA A' B' C'" proof - obtain A0 where P1: "Bet A B A0 \ Cong B A0 A B" using segment_construction by blast moreover obtain A0' where P2: "Bet A' B' A0' \ Cong B' A0' A' B'" using segment_construction by blast moreover have "A0 B C CongA A0' B' C'" proof - have "Acute A0 B C" using assms(1) bet_obtuse__acute P1 cong_diff_3 obtuse_distincts by blast moreover have "A0 B C A0 B C SumA D E F" using P1 acute_distincts assms(2) bet2_suma__suma calculation by blast moreover have "Acute A0' B' C'" using P2 assms(3) bet_obtuse__acute cong_diff_3 obtuse_distincts by blast moreover have "A0' B' C' A0' B' C' SumA D E F" by (metis P2 assms(4) bet2_suma__suma cong_diff_3) ultimately show ?thesis using acute2_suma2__conga by blast qed moreover have "Bet A0 B A" using Bet_perm calculation(1) by blast moreover have "Bet A0' B' A'" using Bet_perm calculation(2) by blast moreover have "A \ B" using assms(1) obtuse_distincts by blast moreover have "A' \ B'" using assms(3) obtuse_distincts by blast ultimately show ?thesis using l11_13 by blast qed lemma bet_suma2__or_conga: assumes "A0 \ B" and "Bet A B A0" and "A B C A B C SumA D E F" and "A' B' C' A' B' C' SumA D E F" shows "A B C CongA A' B' C' \ A0 B C CongA A' B' C'" proof - { fix A' B' C' assume"Acute A' B' C' \ A' B' C' A' B' C' SumA D E F" have "Per A B C \ Acute A B C \ Obtuse A B C" by (metis angle_partition l8_20_1_R1 l8_5) { assume "Per A B C" then have "Bet D E F" using per2_suma__bet assms(3) by auto then have "False" using \Acute A' B' C' \ A' B' C' A' B' C' SumA D E F\ acute_suma__nbet by blast then have "A B C CongA A' B' C' \ A0 B C CongA A' B' C'" by blast } { assume "Acute A B C" have "Acute A' B' C'" by (simp add: \Acute A' B' C' \ A' B' C' A' B' C' SumA D E F\) moreover have "A' B' C' A' B' C' SumA D E F" by (simp add: \Acute A' B' C' \ A' B' C' A' B' C' SumA D E F\) ultimately have "A B C CongA A' B' C' \ A0 B C CongA A' B' C'" using assms(3) \Acute A B C\ acute2_suma2__conga by auto } { assume "Obtuse A B C" have "Acute A0 B C" using \Obtuse A B C\ assms(1) assms(2) bet_obtuse__acute by auto moreover have "A0 B C A0 B C SumA D E F" using assms(1) assms(2) assms(3) bet2_suma__suma by auto ultimately have "A0 B C CongA A' B' C'" using \Acute A' B' C' \ A' B' C' A' B' C' SumA D E F\ acute2_suma2__conga by auto then have "A B C CongA A' B' C' \ A0 B C CongA A' B' C'" by blast } then have "A B C CongA A' B' C' \ A0 B C CongA A' B' C'" using \Acute A B C \ A B C CongA A' B' C' \ A0 B C CongA A' B' C'\ \Per A B C \ A B C CongA A' B' C' \ A0 B C CongA A' B' C'\ \Per A B C \ Acute A B C \ Obtuse A B C\ by blast } then have P1: "\ A' B' C'. (Acute A' B' C' \ A' B' C' A' B' C' SumA D E F) \ (A B C CongA A' B' C' \ A0 B C CongA A' B' C')" by blast have "Per A' B' C' \ Acute A' B' C' \ Obtuse A' B' C'" by (metis angle_partition l8_20_1_R1 l8_5) { assume P2: "Per A' B' C'" have "A B C CongA A' B' C'" proof - have "A \ B \ B \ C" using assms(3) suma_distincts by blast moreover have "A' \ B' \ B' \ C'" using assms(4) suma_distincts by auto moreover have "Per A B C" proof - have "Bet D E F" using P2 assms(4) per2_suma__bet by blast moreover have "A B C A B C SumA D E F" by (simp add: assms(3)) ultimately show ?thesis using assms(3) bet_suma__per by blast qed ultimately show ?thesis using P2 l11_16 by blast qed then have "A B C CongA A' B' C' \ A0 B C CongA A' B' C'" by blast } { assume "Acute A' B' C'" then have "A B C CongA A' B' C' \ A0 B C CongA A' B' C'" using P1 assms(4) by blast } { assume "Obtuse A' B' C'" obtain A0' where "Bet A' B' A0' \ Cong B' A0' A' B'" using segment_construction by blast moreover have "Acute A0' B' C'" using \Obtuse A' B' C'\ bet_obtuse__acute calculation cong_diff_3 obtuse_distincts by blast moreover have "A0' B' C' A0' B' C' SumA D E F" using acute_distincts assms(4) bet2_suma__suma calculation(1) calculation(2) by blast ultimately have "A B C CongA A' B' C' \ A0 B C CongA A' B' C'" using P1 by (metis assms(1) assms(2) assms(3) assms(4) between_symmetry l11_13 suma_distincts) } thus ?thesis using \Acute A' B' C' \ A B C CongA A' B' C' \ A0 B C CongA A' B' C'\ \Per A' B' C' \ A B C CongA A' B' C' \ A0 B C CongA A' B' C'\ \Per A' B' C' \ Acute A' B' C' \ Obtuse A' B' C'\ by blast qed lemma suma2__or_conga_suppa: assumes "A B C A B C SumA D E F" and "A' B' C' A' B' C' SumA D E F" shows "A B C CongA A' B' C' \ A B C SuppA A' B' C'" proof - obtain A0 where P1: "Bet A B A0 \ Cong B A0 A B" using segment_construction by blast then have "A0 \ B" using assms(1) bet_cong_eq suma_distincts by blast then have "A B C CongA A' B' C' \ A0 B C CongA A' B' C'" using assms(1) assms(2) P1 bet_suma2__or_conga by blast thus ?thesis by (metis P1 SuppA_def cong_diff conga_right_comm conga_sym) qed lemma ex_trisuma: assumes "A \ B" and "B \ C" and "A \ C" shows "\ D E F. A B C TriSumA D E F" proof - obtain G H I where "A B C B C A SumA G H I" using assms(1) assms(2) assms(3) ex_suma by presburger moreover then obtain D E F where "G H I C A B SumA D E F" using ex_suma suma_distincts by presburger ultimately show ?thesis using TriSumA_def by blast qed lemma trisuma_perm_231: assumes "A B C TriSumA D E F" shows "B C A TriSumA D E F" proof - obtain A1 B1 C1 where P1: "A B C B C A SumA A1 B1 C1 \ A1 B1 C1 C A B SumA D E F" using TriSumA_def assms(1) by auto obtain A2 B2 C2 where P2: "B C A C A B SumA B2 C2 A2" using P1 ex_suma suma_distincts by fastforce have "A B C B2 C2 A2 SumA D E F" proof - have "SAMS A B C B C A" using assms sams123231 trisuma_distincts by auto moreover have "SAMS B C A C A B" using P2 sams123231 suma_distincts by fastforce ultimately show ?thesis using P1 P2 suma_assoc by blast qed thus ?thesis using P2 TriSumA_def suma_sym by blast qed lemma trisuma_perm_312: assumes "A B C TriSumA D E F" shows "C A B TriSumA D E F" by (simp add: assms trisuma_perm_231) lemma trisuma_perm_321: assumes "A B C TriSumA D E F" shows "C B A TriSumA D E F" proof - obtain G H I where "A B C B C A SumA G H I \ G H I C A B SumA D E F" using TriSumA_def assms(1) by auto thus ?thesis by (meson TriSumA_def suma_comm suma_right_comm suma_sym trisuma_perm_231) qed lemma trisuma_perm_213: assumes "A B C TriSumA D E F" shows "B A C TriSumA D E F" using assms trisuma_perm_231 trisuma_perm_321 by blast lemma trisuma_perm_132: assumes "A B C TriSumA D E F" shows "A C B TriSumA D E F" using assms trisuma_perm_312 trisuma_perm_321 by blast lemma conga_trisuma__trisuma: assumes "A B C TriSumA D E F" and "D E F CongA D' E' F'" shows "A B C TriSumA D' E' F'" proof - obtain G H I where P1: "A B C B C A SumA G H I \ G H I C A B SumA D E F" using TriSumA_def assms(1) by auto have "G H I C A B SumA D' E' F'" proof - have f1: "B \ A" by (metis P1 suma_distincts) have f2: "C \ A" using P1 suma_distincts by blast have "G H I CongA G H I" by (metis (full_types) P1 conga_refl suma_distincts) then show ?thesis using f2 f1 by (meson P1 assms(2) conga3_suma__suma conga_refl) qed thus ?thesis using P1 TriSumA_def by blast qed lemma trisuma2__conga: assumes "A B C TriSumA D E F" and "A B C TriSumA D' E' F'" shows "D E F CongA D' E' F'" proof - obtain G H I where P1: "A B C B C A SumA G H I \ G H I C A B SumA D E F" using TriSumA_def assms(1) by auto then have P1A: "G H I C A B SumA D E F" by simp obtain G' H' I' where P2: "A B C B C A SumA G' H' I' \ G' H' I' C A B SumA D' E' F'" using TriSumA_def assms(2) by auto have "G' H' I' C A B SumA D E F" proof - have "G H I CongA G' H' I'" using P1 P2 suma2__conga by blast moreover have "D E F CongA D E F \ C A B CongA C A B" by (metis assms(1) conga_refl trisuma_distincts) ultimately show ?thesis by (meson P1 conga3_suma__suma) qed thus ?thesis using P2 suma2__conga by auto qed lemma conga3_trisuma__trisuma: assumes "A B C TriSumA D E F" and "A B C CongA A' B' C'" and "B C A CongA B' C' A'" and "C A B CongA C' A' B'" shows "A' B' C' TriSumA D E F" proof - obtain G H I where P1: "A B C B C A SumA G H I \ G H I C A B SumA D E F" using TriSumA_def assms(1) by auto thus ?thesis proof - have "A' B' C' B' C' A' SumA G H I" using conga3_suma__suma P1 by (meson assms(2) assms(3) suma2__conga) moreover have "G H I C' A' B' SumA D E F" using conga3_suma__suma P1 by (meson P1 assms(4) suma2__conga) ultimately show ?thesis using TriSumA_def by blast qed qed lemma col_trisuma__bet: assumes "Col A B C" and "A B C TriSumA P Q R" shows "Bet P Q R" proof - obtain D E F where P1: "A B C B C A SumA D E F \ D E F C A B SumA P Q R" using TriSumA_def assms(2) by auto { assume "Bet A B C" have "A B C CongA P Q R" proof - have "A B C CongA D E F" proof - have "C \ A \ C \ B" using assms(2) trisuma_distincts by blast then have "C Out B A" using \ Bet A B C\ bet_out_1 by fastforce thus ?thesis using P1 out546_suma__conga by auto qed moreover have "D E F CongA P Q R" proof - have "A \ C \ A \ B" using assms(2) trisuma_distincts by blast then have "A Out C B" using Out_def \Bet A B C\ by auto thus ?thesis using P1 out546_suma__conga by auto qed ultimately show ?thesis using conga_trans by blast qed then have "Bet P Q R" using \Bet A B C\ bet_conga__bet by blast } { assume "Bet B C A" have "B C A CongA P Q R" proof - have "B C A CongA D E F" proof - have "B \ A \ B \ C" using assms(2) trisuma_distincts by blast then have "B Out A C" using Out_def \Bet B C A\ by auto thus ?thesis using P1 out213_suma__conga by blast qed moreover have "D E F CongA P Q R" proof - have "A \ C \ A \ B" using assms(2) trisuma_distincts by auto then have "A Out C B" using \Bet B C A\ bet_out_1 by auto thus ?thesis using P1 out546_suma__conga by blast qed ultimately show ?thesis using not_conga by blast qed then have "Bet P Q R" using \Bet B C A\ bet_conga__bet by blast } { assume "Bet C A B" have "E Out D F" proof - have "C Out B A" using \Bet C A B\ assms(2) bet_out l6_6 trisuma_distincts by blast moreover have "B C A CongA D E F" proof - have "B \ A \ B \ C" using P1 suma_distincts by blast then have "B Out A C" using \Bet C A B\ bet_out_1 by auto thus ?thesis using out213_suma__conga P1 by blast qed ultimately show ?thesis using l11_21_a by blast qed then have "C A B CongA P Q R" using P1 out213_suma__conga by blast then have "Bet P Q R" using \Bet C A B\ bet_conga__bet by blast } thus ?thesis using Col_def \Bet A B C \ Bet P Q R\ \Bet B C A \ Bet P Q R\ assms(1) by blast qed lemma suma_dec: "A B C D E F SumA G H I \ \ A B C D E F SumA G H I" by simp lemma sams_dec: "SAMS A B C D E F \ \ SAMS A B C D E F" by simp lemma trisuma_dec: "A B C TriSumA P Q R \ \ A B C TriSumA P Q R" by simp subsection "Parallelism" lemma par_reflexivity: assumes "A \ B" shows "A B Par A B" using Par_def assms not_col_distincts by blast lemma par_strict_irreflexivity: "\ A B ParStrict A B" using ParStrict_def col_trivial_3 by blast lemma not_par_strict_id: "\ A B ParStrict A C" using ParStrict_def col_trivial_1 by blast lemma par_id: assumes "A B Par A C" shows "Col A B C" using Col_cases Par_def assms not_par_strict_id by auto lemma par_strict_not_col_1: assumes "A B ParStrict C D" shows "\ Col A B C" using Col_perm ParStrict_def assms col_trivial_1 by blast lemma par_strict_not_col_2: assumes "A B ParStrict C D" shows "\ Col B C D" using ParStrict_def assms col_trivial_3 by auto lemma par_strict_not_col_3: assumes "A B ParStrict C D" shows "\ Col C D A" using Col_perm ParStrict_def assms col_trivial_1 by blast lemma par_strict_not_col_4: assumes "A B ParStrict C D" shows "\ Col A B D" using Col_perm ParStrict_def assms col_trivial_3 by blast lemma par_id_1: assumes "A B Par A C" shows "Col B A C" using Par_def assms not_par_strict_id by auto lemma par_id_2: assumes "A B Par A C" shows "Col B C A" using Col_perm assms par_id_1 by blast lemma par_id_3: assumes "A B Par A C" shows "Col A C B" using Col_perm assms par_id_2 by blast lemma par_id_4: assumes "A B Par A C" shows "Col C B A" using Col_perm assms par_id_2 by blast lemma par_id_5: assumes "A B Par A C" shows "Col C A B" using assms col_permutation_2 par_id by blast lemma par_strict_symmetry: assumes "A B ParStrict C D" shows "C D ParStrict A B" using ParStrict_def assms coplanar_perm_16 by blast lemma par_symmetry: assumes "A B Par C D" shows "C D Par A B" by (smt NCol_perm Par_def assms l6_16_1 par_strict_symmetry) lemma par_left_comm: assumes "A B Par C D" shows "B A Par C D" by (metis (mono_tags, lifting) ParStrict_def Par_def assms ncoplanar_perm_6 not_col_permutation_5) lemma par_right_comm: assumes "A B Par C D" shows "A B Par D C" using assms par_left_comm par_symmetry by blast lemma par_comm: assumes "A B Par C D" shows "B A Par D C" using assms par_left_comm par_right_comm by blast lemma par_strict_left_comm: assumes "A B ParStrict C D" shows "B A ParStrict C D" using ParStrict_def assms ncoplanar_perm_6 not_col_permutation_5 by blast lemma par_strict_right_comm: assumes "A B ParStrict C D" shows "A B ParStrict D C" using assms par_strict_left_comm par_strict_symmetry by blast lemma par_strict_comm: assumes "A B ParStrict C D" shows "B A ParStrict D C" by (simp add: assms par_strict_left_comm par_strict_right_comm) lemma par_strict_neq1: assumes "A B ParStrict C D" shows "A \ B" using assms col_trivial_1 par_strict_not_col_4 by blast lemma par_strict_neq2: assumes "A B ParStrict C D" shows "C \ D" using assms col_trivial_2 par_strict_not_col_2 by blast lemma par_neq1: assumes "A B Par C D" shows "A \ B" using Par_def assms par_strict_neq1 by blast lemma par_neq2: assumes "A B Par C D" shows "C \ D" using assms par_neq1 par_symmetry by blast lemma Par_cases: assumes "A B Par C D \ B A Par C D \ A B Par D C \ B A Par D C \ C D Par A B \ C D Par B A \ D C Par A B \ D C Par B A" shows "A B Par C D" using assms par_right_comm par_symmetry by blast lemma Par_perm: assumes "A B Par C D" shows "A B Par C D \ B A Par C D \ A B Par D C \ B A Par D C \ C D Par A B \ C D Par B A \ D C Par A B \ D C Par B A" using Par_cases assms by blast lemma Par_strict_cases: assumes "A B ParStrict C D \ B A ParStrict C D \ A B ParStrict D C \ B A ParStrict D C \ C D ParStrict A B \ C D ParStrict B A \ D C ParStrict A B \ D C ParStrict B A" shows "A B ParStrict C D" using assms par_strict_right_comm par_strict_symmetry by blast lemma Par_strict_perm: assumes "A B ParStrict C D" shows "A B ParStrict C D \ B A ParStrict C D \ A B ParStrict D C \ B A ParStrict D C \ C D ParStrict A B \ C D ParStrict B A \ D C ParStrict A B \ D C ParStrict B A" using Par_strict_cases assms by blast lemma l12_6: assumes "A B ParStrict C D" shows "A B OS C D" by (metis Col_def ParStrict_def Par_strict_perm TS_def assms cop_nts__os par_strict_not_col_2) lemma pars__os3412: assumes "A B ParStrict C D" shows "C D OS A B" by (simp add: assms l12_6 par_strict_symmetry) lemma perp_dec: "A B Perp C D \ \ A B Perp C D" by simp lemma col_cop2_perp2__col: assumes "X1 X2 Perp A B" and "Y1 Y2 Perp A B" and "Col X1 Y1 Y2" and "Coplanar A B X2 Y1" and "Coplanar A B X2 Y2" shows "Col X2 Y1 Y2" proof cases assume "X1 = Y2" thus ?thesis using assms(1) assms(2) assms(4) cop_perp2__col not_col_permutation_2 perp_left_comm by blast next assume "X1 \ Y2" then have "Y2 X1 Perp A B" by (metis Col_cases assms(2) assms(3) perp_col perp_comm perp_right_comm) then have P1: "X1 Y2 Perp A B" using Perp_perm by blast thus ?thesis proof cases assume "X1 = Y1" thus ?thesis using assms(1) assms(2) assms(5) cop_perp2__col not_col_permutation_4 by blast next assume "X1 \ Y1" then have "X1 Y1 Perp A B" using Col_cases P1 assms(3) perp_col by blast thus ?thesis using P1 assms(1) assms(4) assms(5) col_transitivity_2 cop_perp2__col perp_not_eq_1 by blast qed qed lemma col_perp2_ncol_col: assumes "X1 X2 Perp A B" and "Y1 Y2 Perp A B" and "Col X1 Y1 Y2" and "\ Col X1 A B" shows "Col X2 Y1 Y2" proof - have "Coplanar A B X2 Y1" proof cases assume "X1 = Y1" thus ?thesis using assms(1) ncoplanar_perm_22 perp__coplanar by blast next assume "X1 \ Y1" then have "Y1 X1 Perp A B" by (metis Col_cases assms(2) assms(3) perp_col) thus ?thesis by (meson assms(1) assms(4) coplanar_trans_1 ncoplanar_perm_18 ncoplanar_perm_4 perp__coplanar) qed then moreover have "Coplanar A B X2 Y2" by (smt assms(1) assms(2) assms(3) assms(4) col_cop2__cop coplanar_perm_17 coplanar_perm_18 coplanar_trans_1 perp__coplanar) ultimately show ?thesis using assms(1) assms(2) assms(3) col_cop2_perp2__col by blast qed lemma l12_9: assumes "Coplanar C1 C2 A1 B1" and "Coplanar C1 C2 A1 B2" and "Coplanar C1 C2 A2 B1" and "Coplanar C1 C2 A2 B2" and "A1 A2 Perp C1 C2" and "B1 B2 Perp C1 C2" shows "A1 A2 Par B1 B2" proof - have P1: "A1 \ A2 \ C1 \ C2" using assms(5) perp_distinct by auto have P2: "B1 \ B2" using assms(6) perp_distinct by auto show ?thesis proof cases assume "Col A1 B1 B2" then show ?thesis using P1 P2 Par_def assms(3) assms(4) assms(5) assms(6) col_cop2_perp2__col by blast next assume P3: "\ Col A1 B1 B2" { assume "\ Col C1 C2 A1" then have "Coplanar A1 A2 B1 B2" by (smt assms(1) assms(2) assms(5) coplanar_perm_22 coplanar_perm_8 coplanar_pseudo_trans ncop_distincts perp__coplanar) } have "C1 C2 Perp A1 A2" using Perp_cases assms(5) by blast then have "Coplanar A1 A2 B1 B2" by (smt \\ Col C1 C2 A1 \ Coplanar A1 A2 B1 B2\ assms(3) assms(4) coplanar_perm_1 coplanar_pseudo_trans ncop_distincts perp__coplanar perp_not_col2) { assume "\ X. Col X A1 A2 \ Col X B1 B2" then obtain AB where P4: "Col AB A1 A2 \ Col AB B1 B2" by auto then have "False" proof cases assume "AB = A1" thus ?thesis using P3 P4 by blast next assume "AB \ A1" then have "A1 AB Perp C1 C2" by (metis P4 assms(5) not_col_permutation_2 perp_col) then have "AB A1 Perp C1 C2" by (simp add: perp_left_comm) thus ?thesis using P3 P4 assms(1) assms(2) assms(6) col_cop2_perp2__col by blast qed } then show ?thesis using ParStrict_def Par_def \Coplanar A1 A2 B1 B2\ by blast qed qed lemma parallel_existence: assumes "A \ B" shows "\ C D. C \ D \ A B Par C D \ Col P C D" proof cases assume "Col A B P" then show ?thesis using Col_perm assms par_reflexivity by blast next assume P1: "\ Col A B P" then obtain P' where P2: "Col A B P' \ A B Perp P P'" using l8_18_existence by blast then have P3: "P \ P'" using P1 by blast show ?thesis proof cases assume P4: "P' = A" have "\ Q. Per Q P A \ Cong Q P A B \ A P OS Q B" proof - have "Col A P P" using not_col_distincts by auto moreover have "\ Col A P B" by (simp add: P1 not_col_permutation_5) ultimately show ?thesis using P3 P4 assms ex_per_cong by simp qed then obtain Q where T1: "Per Q P A \ Cong Q P A B \ A P OS Q B" by auto then have T2: "P \ Q" using os_distincts by auto have T3: "A B Par P Q" proof - have "P Q Perp P A" proof - have "P \ A" using P3 P4 by auto moreover have "Col P P Q" by (simp add: col_trivial_1) ultimately show ?thesis by (metis T1 T2 Tarski_neutral_dimensionless.Perp_perm Tarski_neutral_dimensionless_axioms per_perp) qed moreover have "Coplanar P A A P" using ncop_distincts by auto moreover have "Coplanar P A B P" using ncop_distincts by auto moreover have "Coplanar P A B Q" by (metis (no_types) T1 ncoplanar_perm_7 os__coplanar) moreover have "A B Perp P A" using P2 P4 by auto ultimately show ?thesis using l12_9 ncop_distincts by blast qed thus ?thesis using T2 col_trivial_1 by auto next assume T4: "P' \ A" have "\ Q. Per Q P P' \ Cong Q P A B \ P' P OS Q A" proof - have "P' \ P" using P3 by auto moreover have "A \ B" by (simp add: assms) moreover have "Col P' P P" using not_col_distincts by blast moreover have "\ Col P' P A" by (metis P1 P2 T4 col2__eq col_permutation_1) ultimately show ?thesis using ex_per_cong by blast qed then obtain Q where T5: "Per Q P P' \ Cong Q P A B \ P' P OS Q A" by blast then have T6: "P \ Q" using os_distincts by blast moreover have "A B Par P Q" proof - have "Coplanar P P' A P" using ncop_distincts by auto moreover have "Coplanar P P' A Q" by (meson T5 ncoplanar_perm_7 os__coplanar) then moreover have "Coplanar P P' B Q" by (smt P2 T4 col2_cop__cop col_permutation_5 col_transitivity_1 coplanar_perm_5) moreover have "Coplanar P P' B P" using ncop_distincts by auto moreover have "A B Perp P P'" by (simp add: P2) moreover have "P Q Perp P P'" by (metis P3 T5 T6 Tarski_neutral_dimensionless.Perp_perm Tarski_neutral_dimensionless_axioms per_perp) ultimately show ?thesis using l12_9 by blast qed moreover have "Col P P Q" by (simp add: col_trivial_1) ultimately show ?thesis by blast qed qed lemma par_col_par: assumes "C \ D'" and "A B Par C D" and "Col C D D'" shows "A B Par C D'" proof - { assume P1: "A B ParStrict C D" have "Coplanar A B C D'" using assms(2) assms(3) col2__eq col2_cop__cop par__coplanar par_neq2 by blast then have "A B Par C D'" by (smt ParStrict_def Par_def P1 assms(1) assms(3) colx not_col_distincts not_col_permutation_5) } { assume "A \ B \ C \ D \ Col A C D \ Col B C D" then have "A B Par C D'" using Par_def assms(1) assms(3) col2__eq col_permutation_2 by blast } thus ?thesis using Par_def \A B ParStrict C D \ A B Par C D'\ assms(2) by auto qed lemma parallel_existence1: assumes "A \ B" shows "\ Q. A B Par P Q" proof - obtain C D where "C \ D \ A B Par C D \ Col P C D" using assms parallel_existence by blast then show ?thesis by (metis Col_cases Par_cases par_col_par) qed lemma par_not_col: assumes "A B ParStrict C D" and "Col X A B" shows "\ Col X C D" using ParStrict_def assms(1) assms(2) by blast lemma not_strict_par1: assumes "A B Par C D" and "Col A B X" and "Col C D X" shows "Col A B C" by (smt Par_def assms(1) assms(2) assms(3) col2__eq col_permutation_2 par_not_col) lemma not_strict_par2: assumes "A B Par C D" and "Col A B X" and "Col C D X" shows "Col A B D" using Par_cases assms(1) assms(2) assms(3) not_col_permutation_4 not_strict_par1 by blast lemma not_strict_par: assumes "A B Par C D" and "Col A B X" and "Col C D X" shows "Col A B C \ Col A B D" using assms(1) assms(2) assms(3) not_strict_par1 not_strict_par2 by blast lemma not_par_not_col: assumes "A \ B" and "A \ C" and "\ A B Par A C" shows "\ Col A B C" using Par_def assms(1) assms(2) assms(3) not_col_distincts not_col_permutation_4 by blast lemma not_par_inter_uniqueness: assumes "A \ B" and "C \ D" and "\ A B Par C D" and "Col A B X" and "Col C D X" and "Col A B Y" and "Col C D Y" shows "X = Y" proof cases assume P1: "C = Y" thus ?thesis proof cases assume P2: "C = X" thus ?thesis using P1 by auto next assume "C \ X" thus ?thesis by (smt Par_def assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) col3 col_permutation_5 l6_21) qed next assume "C \ Y" thus ?thesis by (smt Par_def assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) col_permutation_2 col_permutation_4 l6_21) qed lemma inter_uniqueness_not_par: assumes "\ Col A B C" and "Col A B P" and "Col C D P" shows "\ A B Par C D" using assms(1) assms(2) assms(3) not_strict_par1 by blast lemma col_not_col_not_par: assumes "\ P. Col A B P \ Col C D P" and "\ Q. Col C D Q \ \Col A B Q" shows "\ A B Par C D" using assms(1) assms(2) colx not_strict_par par_neq2 by blast lemma par_distincts: assumes "A B Par C D" shows "A B Par C D \ A \ B \ C \ D" using assms par_neq1 par_neq2 by blast lemma par_not_col_strict: assumes "A B Par C D" and "Col C D P" and "\ Col A B P" shows "A B ParStrict C D" using Col_cases Par_def assms(1) assms(2) assms(3) col3 by blast lemma col_cop_perp2_pars: assumes "\ Col A B P" and "Col C D P" and "Coplanar A B C D" and "A B Perp P Q" and "C D Perp P Q" shows "A B ParStrict C D" proof - have P1: "C \ D" using assms(5) perp_not_eq_1 by auto then have P2: "Coplanar A B C P" using col_cop__cop assms(2) assms(3) by blast moreover have P3: "Coplanar A B D P" using col_cop__cop using P1 assms(2) assms(3) col2_cop__cop col_trivial_2 by blast have "A B Par C D" proof - have "Coplanar P A Q C" proof - have "\ Col B P A" by (simp add: assms(1) not_col_permutation_1) moreover have "Coplanar B P A Q" by (meson assms(4) ncoplanar_perm_12 perp__coplanar) moreover have "Coplanar B P A C" using P2 ncoplanar_perm_13 by blast ultimately show ?thesis using coplanar_trans_1 by auto qed then have P4: "Coplanar P Q A C" using ncoplanar_perm_2 by blast have "Coplanar P A Q D" proof - have "\ Col B P A" by (simp add: assms(1) not_col_permutation_1) moreover have "Coplanar B P A Q" by (meson assms(4) ncoplanar_perm_12 perp__coplanar) moreover have "Coplanar B P A D" using P3 ncoplanar_perm_13 by blast ultimately show ?thesis using coplanar_trans_1 by blast qed then moreover have "Coplanar P Q A D" using ncoplanar_perm_2 by blast moreover have "Coplanar P Q B C" using P2 assms(1) assms(4) coplanar_perm_1 coplanar_perm_10 coplanar_trans_1 perp__coplanar by blast moreover have "Coplanar P Q B D" by (meson P3 assms(1) assms(4) coplanar_trans_1 ncoplanar_perm_1 ncoplanar_perm_13 perp__coplanar) ultimately show ?thesis using assms(4) assms(5) l12_9 P4 by auto qed thus ?thesis using assms(1) assms(2) par_not_col_strict by auto qed lemma all_one_side_par_strict: assumes "C \ D" and "\ P. Col C D P \ A B OS C P" shows "A B ParStrict C D" proof - have P1: "Coplanar A B C D" by (simp add: assms(2) col_trivial_2 os__coplanar) { assume "\ X. Col X A B \ Col X C D" then obtain X where P2: "Col X A B \ Col X C D" by blast have "A B OS C X" by (simp add: P2 Col_perm assms(2)) then obtain M where "A B TS C M \ A B TS X M" by (meson Col_cases P2 col124__nos) then have "False" using P2 TS_def by blast } thus ?thesis using P1 ParStrict_def by auto qed lemma par_col_par_2: assumes "A \ P" and "Col A B P" and "A B Par C D" shows "A P Par C D" using assms(1) assms(2) assms(3) par_col_par par_symmetry by blast lemma par_col2_par: assumes "E \ F" and "A B Par C D" and "Col C D E" and "Col C D F" shows "A B Par E F" by (metis assms(1) assms(2) assms(3) assms(4) col_transitivity_2 not_col_permutation_4 par_col_par par_distincts par_right_comm) lemma par_col2_par_bis: assumes "E \ F" and "A B Par C D" and "Col E F C" and "Col E F D" shows "A B Par E F" by (metis assms(1) assms(2) assms(3) assms(4) col_transitivity_1 not_col_permutation_2 par_col2_par) lemma par_strict_col_par_strict: assumes "C \ E" and "A B ParStrict C D" and "Col C D E" shows "A B ParStrict C E" proof - have P1: "C E Par A B" using Par_def Par_perm assms(1) assms(2) assms(3) par_col_par_2 by blast { assume "C E ParStrict A B" then have "A B ParStrict C E" by (metis par_strict_symmetry) } thus ?thesis using Col_cases Par_def P1 assms(2) par_strict_not_col_1 by blast qed lemma par_strict_col2_par_strict: assumes "E \ F" and "A B ParStrict C D" and "Col C D E" and "Col C D F" shows "A B ParStrict E F" by (smt ParStrict_def assms(1) assms(2) assms(3) assms(4) col2_cop__cop colx not_col_permutation_1 par_strict_neq1 par_strict_symmetry) lemma line_dec: "(Col C1 B1 B2 \ Col C2 B1 B2) \ \ (Col C1 B1 B2 \ Col C2 B1 B2)" by simp lemma par_distinct: assumes "A B Par C D" shows "A \ B \ C \ D" using assms par_neq1 par_neq2 by auto lemma par_col4__par: assumes "E \ F" and "G \ H" and "A B Par C D" and "Col A B E" and "Col A B F" and "Col C D G" and "Col C D H" shows "E F Par G H" proof - have "C D Par E F" using Par_cases assms(1) assms(3) assms(4) assms(5) par_col2_par by blast then have "E F Par C D" by (simp add: \C D Par E F\ par_symmetry) thus ?thesis using assms(2) assms(6) assms(7) par_col2_par by blast qed lemma par_strict_col4__par_strict: assumes "E \ F" and "G \ H" and "A B ParStrict C D" and "Col A B E" and "Col A B F" and "Col C D G" and "Col C D H" shows "E F ParStrict G H" proof - have "C D ParStrict E F" using Par_strict_cases assms(1) assms(3) assms(4) assms(5) par_strict_col2_par_strict by blast then have "E F ParStrict C D" by (simp add: \C D ParStrict E F\ par_strict_symmetry) thus ?thesis using assms(2) assms(6) assms(7) par_strict_col2_par_strict by blast qed lemma par_strict_one_side: assumes "A B ParStrict C D" and "Col C D P" shows "A B OS C P" proof cases assume "C = P" thus ?thesis using assms(1) assms(2) not_col_permutation_5 one_side_reflexivity par_not_col by blast next assume "C \ P" thus ?thesis using assms(1) assms(2) l12_6 par_strict_col_par_strict by blast qed lemma par_strict_all_one_side: assumes "A B ParStrict C D" shows "\ P. Col C D P \ A B OS C P" using assms par_strict_one_side by blast lemma inter_trivial: assumes "\ Col A B X" shows "X Inter A X B X" by (metis Col_perm Inter_def assms col_trivial_1) lemma inter_sym: assumes "X Inter A B C D" shows "X Inter C D A B" proof - obtain P where P1: "Col P C D \ \ Col P A B" using Inter_def assms by auto have P2: "A \ B" using P1 col_trivial_2 by blast then show ?thesis proof cases assume "A = X" have "Col B A B" by (simp add: col_trivial_3) { assume P3: "Col B C D" have "Col P A B" proof - have "C \ D" using Inter_def assms by blast moreover have "Col C D P" using P1 not_col_permutation_2 by blast moreover have "Col C D A" using Inter_def \A = X\ assms by auto moreover have "Col C D B" using P3 not_col_permutation_2 by blast ultimately show ?thesis using col3 by blast qed then have "False" by (simp add: P1) } then have "\ Col B C D" by auto then show ?thesis using Inter_def P2 assms by (meson col_trivial_3) next assume P5: "A \ X" have P6: "Col A A B" using not_col_distincts by blast { assume P7: "Col A C D" have "Col A P X" proof - have "C \ D" using Inter_def assms by auto moreover have "Col C D A" using Col_cases P7 by blast moreover have "Col C D P" using Col_cases P1 by auto moreover have "Col C D X" using Inter_def assms by auto ultimately show ?thesis using col3 by blast qed then have "Col P A B" by (metis (full_types) Col_perm Inter_def P5 assms col_transitivity_2) then have "False" by (simp add: P1) } then have "\ Col A C D" by auto then show ?thesis by (meson Inter_def P2 assms col_trivial_1) qed qed lemma inter_left_comm: assumes "X Inter A B C D" shows "X Inter B A C D" using Col_cases Inter_def assms by auto lemma inter_right_comm: assumes "X Inter A B C D" shows "X Inter A B D C" by (metis assms inter_left_comm inter_sym) lemma inter_comm: assumes "X Inter A B C D" shows "X Inter B A D C" using assms inter_left_comm inter_right_comm by blast lemma l12_17: assumes "A \ B" and "P Midpoint A C" and "P Midpoint B D" shows "A B Par C D" proof cases assume P1: "Col A B P" thus ?thesis proof cases assume "A = P" thus ?thesis using assms(1) assms(2) assms(3) cong_diff_2 is_midpoint_id midpoint_col midpoint_cong not_par_not_col by blast next assume P2: "A \ P" thus ?thesis proof cases assume "B = P" thus ?thesis by (metis assms(1) assms(2) assms(3) midpoint_col midpoint_distinct_2 midpoint_distinct_3 not_par_not_col par_comm) next assume P3: "B \ P" have P4: "Col B P D" using assms(3) midpoint_col not_col_permutation_4 by blast have P5: "Col A P C" using assms(2) midpoint_col not_col_permutation_4 by blast then have P6: "Col B C P" using P1 P2 col_transitivity_2 not_col_permutation_3 not_col_permutation_5 by blast have "C \ D" using assms(1) assms(2) assms(3) l7_9 by blast moreover have "Col A C D" using P1 P3 P4 P6 col3 not_col_permutation_3 not_col_permutation_5 by blast moreover have "Col B C D" using P3 P4 P6 col_trivial_3 colx by blast ultimately show ?thesis by (simp add: Par_def assms(1)) qed qed next assume T1: "\ Col A B P" then obtain E where T2: "Col A B E \ A B Perp P E" using l8_18_existence by blast have T3: "A \ P" using T1 col_trivial_3 by blast then show ?thesis proof cases assume T4: "A = E" then have T5: "Per P A B" using T2 l8_2 perp_per_1 by blast obtain B' where T6: "Bet B A B' \ Cong A B' B A" using segment_construction by blast obtain D' where T7: "Bet B' P D' \ Cong P D' B' P" using segment_construction by blast have T8: "C Midpoint D D'" using T6 T7 assms(2) assms(3) midpoint_def not_cong_3412 symmetry_preserves_midpoint by blast have "Col A B B'" using Col_cases Col_def T6 by blast then have T9: "Per P A B'" using per_col T5 assms(1) by blast obtain B'' where T10: "A Midpoint B B'' \ Cong P B P B''" using Per_def T5 by auto then have "B' = B''" using T6 cong_symmetry midpoint_def symmetric_point_uniqueness by blast then have "Cong P D P D'" by (metis Cong_perm Midpoint_def T10 T7 assms(3) cong_inner_transitivity) then have T12: "Per P C D" using Per_def T8 by auto then have T13: "C PerpAt P C C D" by (metis T3 assms(1) assms(2) assms(3) l7_3_2 per_perp_in sym_preserve_diff) have T14: "P \ C" using T3 assms(2) is_midpoint_id_2 by auto have T15: "C \ D" using assms(1) assms(2) assms(3) l7_9 by auto have T15A: "C C Perp C D \ P C Perp C D" using T12 T14 T15 per_perp by auto { assume "C C Perp C D" then have "A B Par C D" using perp_distinct by auto } { assume "P C Perp C D" have "A B Par C D" proof - have "Coplanar P A A C" using ncop_distincts by blast moreover have "Coplanar P A A D" using ncop_distincts by blast moreover have "Coplanar P A B C" by (simp add: assms(2) coplanar_perm_1 midpoint__coplanar) moreover have "Coplanar P A B D" using assms(3) midpoint_col ncop__ncols by blast moreover have "A B Perp P A" using T2 T4 by auto moreover have "C D Perp P A" proof - have "P A Perp C D" proof - have "P \ A" using T3 by auto moreover have "P C Perp C D" using T14 T15 T12 per_perp by blast moreover have "Col P C A" by (simp add: assms(2) l7_2 midpoint_col) ultimately show ?thesis using perp_col by blast qed then show ?thesis using Perp_perm by blast qed ultimately show ?thesis using l12_9 by blast qed } then show ?thesis using T15A using \C C Perp C D \ A B Par C D\ by blast next assume S1B: "A \ E" obtain F where S2: "Bet E P F \ Cong P F E P" using segment_construction by blast then have S2A: "P Midpoint E F" using midpoint_def not_cong_3412 by blast then have S3: "Col C D F" using T2 assms(2) assms(3) mid_preserves_col by blast obtain A' where S4: "Bet A E A' \ Cong E A' A E" using segment_construction by blast obtain C' where S5: "Bet A' P C' \ Cong P C' A' P" using segment_construction by blast have S6: "F Midpoint C C'" using S4 S5 S2A assms(2) midpoint_def not_cong_3412 symmetry_preserves_midpoint by blast have S7: "Per P E A" using T2 col_trivial_3 l8_16_1 by blast have S8: "Cong P C P C'" proof - have "Cong P C P A" using Cong_perm Midpoint_def assms(2) by blast moreover have "Cong P A P C'" proof - obtain A'' where S9: "E Midpoint A A'' \ Cong P A P A''" using Per_def S7 by blast have S10: "A' = A''" using Cong_perm Midpoint_def S4 S9 symmetric_point_uniqueness by blast then have "Cong P A P A'" using S9 by auto moreover have "Cong P A' P C'" using Cong_perm S5 by blast ultimately show ?thesis using cong_transitivity by blast qed ultimately show ?thesis using cong_transitivity by blast qed then have S9: "Per P F C" using S6 Per_def by blast then have "F PerpAt P F F C" by (metis S2 S2A T1 T2 S1B assms(2) cong_diff_3 l7_9 per_perp_in) then have "F PerpAt F P C F" using Perp_in_perm by blast then have S10: "F P Perp C F \ F F Perp C F" using l8_15_2 perp_in_col by blast { assume S11: "F P Perp C F" have "Coplanar P E A C" proof - have "Col P E P \ Col A C P" using assms(2) col_trivial_3 midpoint_col not_col_permutation_2 by blast then show ?thesis using Coplanar_def by blast qed moreover have "Coplanar P E A D" proof - have "Col P D B \ Col E A B" using Mid_cases T2 assms(3) midpoint_col not_col_permutation_1 by blast then show ?thesis using Coplanar_def by blast qed moreover have "Coplanar P E B C" by (metis S1B T2 calculation(1) col2_cop__cop col_transitivity_1 ncoplanar_perm_5 not_col_permutation_5) moreover have "Coplanar P E B D" by (metis S1B T2 calculation(2) col2_cop__cop col_transitivity_1 ncoplanar_perm_5 not_col_permutation_5) moreover have "C D Perp P E" proof - have "C \ D" using assms(1) assms(2) assms(3) sym_preserve_diff by blast moreover have "P F Perp C F" using Perp_perm S11 by blast moreover have "Col P F E" by (simp add: Col_def S2) moreover have "Col C F D" using Col_perm S3 by blast ultimately show ?thesis using per_col by (smt Perp_cases S2 col_trivial_3 cong_diff perp_col4 perp_not_eq_1) qed ultimately have "A B Par C D" using T2 l12_9 by blast } { assume "F F Perp C F" then have "A B Par C D" using perp_distinct by blast } thus ?thesis using S10 \F P Perp C F \ A B Par C D\ by blast qed qed lemma l12_18_a: assumes "Cong A B C D" and "Cong B C D A" and "\ Col A B C" and "B \ D" and "Col A P C" and "Col B P D" shows "A B Par C D" proof - have "P Midpoint A C \ P Midpoint B D" using assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) l7_21 by blast then show ?thesis using assms(3) l12_17 not_col_distincts by blast qed lemma l12_18_b: assumes "Cong A B C D" and "Cong B C D A" and "\ Col A B C" and "B \ D" and "Col A P C" and "Col B P D" shows "B C Par D A" by (smt assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) cong_symmetry inter_uniqueness_not_par l12_18_a l6_21 not_col_distincts) lemma l12_18_c: assumes "Cong A B C D" and "Cong B C D A" and "\ Col A B C" and "B \ D" and "Col A P C" and "Col B P D" shows "B D TS A C" proof - have "P Midpoint A C \ P Midpoint B D" using assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) l7_21 by blast then show ?thesis proof - have "A C TS B D" by (metis Col_cases Tarski_neutral_dimensionless.mid_two_sides Tarski_neutral_dimensionless_axioms \P Midpoint A C \ P Midpoint B D\ assms(3)) then have "\ Col B D A" by (meson Col_cases Tarski_neutral_dimensionless.mid_preserves_col Tarski_neutral_dimensionless.ts__ncol Tarski_neutral_dimensionless_axioms \P Midpoint A C \ P Midpoint B D\ l7_2) then show ?thesis by (meson Tarski_neutral_dimensionless.mid_two_sides Tarski_neutral_dimensionless_axioms \P Midpoint A C \ P Midpoint B D\) qed qed lemma l12_18_d: assumes "Cong A B C D" and "Cong B C D A" and "\ Col A B C" and "B \ D" and "Col A P C" and "Col B P D" shows "A C TS B D" by (metis (no_types, lifting) Col_cases TS_def assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) l12_18_c not_col_distincts not_cong_2143 not_cong_4321) lemma l12_18: assumes "Cong A B C D" and "Cong B C D A" and "\ Col A B C" and "B \ D" and "Col A P C" and "Col B P D" shows "A B Par C D \ B C Par D A \ B D TS A C \ A C TS B D" using assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) l12_18_a l12_18_b l12_18_c l12_18_d by auto lemma par_two_sides_two_sides: assumes "A B Par C D" and "B D TS A C" shows "A C TS B D" by (metis Par_def TS_def assms(1) assms(2) invert_one_side invert_two_sides l12_6 l9_31 not_col_permutation_4 one_side_symmetry os_ts1324__os pars__os3412) lemma par_one_or_two_sides: assumes "A B ParStrict C D" shows "(A C TS B D \ B D TS A C) \ (A C OS B D \ B D OS A C)" by (smt Par_def assms invert_one_side l12_6 l9_31 not_col_permutation_3 os_ts1324__os par_strict_not_col_1 par_strict_not_col_2 par_two_sides_two_sides pars__os3412 two_sides_cases) lemma l12_21_b: assumes "A C TS B D" and "B A C CongA D C A" shows "A B Par C D" proof - have P1: "\ Col A B C" using TS_def assms(1) not_col_permutation_4 by blast then have P2: "A \ B" using col_trivial_1 by auto have P3: "C \ D" using assms(1) ts_distincts by blast then obtain D' where P4: "C Out D D' \ Cong C D' A B" using P2 segment_construction_3 by blast have P5: "B A C CongA D' C A" proof - have "A Out B B" using P2 out_trivial by auto moreover have "A Out C C" using P1 col_trivial_3 out_trivial by force moreover have "C Out D' D" by (simp add: P4 l6_6) moreover have "C Out A A" using P1 not_col_distincts out_trivial by auto ultimately show ?thesis using assms(2) l11_10 by blast qed then have P6: "Cong D' A B C" using Cong_perm P4 cong_pseudo_reflexivity l11_49 by blast have P7: "A C TS D' B" proof - have "A C TS D B" by (simp add: assms(1) l9_2) moreover have "Col C A C" using col_trivial_3 by auto ultimately show ?thesis using P4 l9_5 by blast qed then obtain M where P8: "Col M A C \ Bet D' M B" using TS_def by blast have "B \ D'" using P7 not_two_sides_id by blast then have "M Midpoint A C \ M Midpoint B D'" by (metis Col_cases P1 P4 P6 P8 bet_col l7_21 not_cong_3412) then have "A B Par C D'" using P2 l12_17 by blast thus ?thesis by (meson P3 P4 Tarski_neutral_dimensionless.par_col_par Tarski_neutral_dimensionless_axioms l6_6 out_col) qed lemma l12_22_aux: assumes "P \ A" and "A \ C" and "Bet P A C" and "P A OS B D" and "B A P CongA D C P" shows "A B Par C D" proof - have P1: "P \ C" using CongA_def assms(5) by blast obtain B' where P2: "Bet B A B' \ Cong A B' B A" using segment_construction by blast have P3: "P A B CongA C A B'" by (metis CongA_def P2 assms(2) assms(3) assms(5) cong_reverse_identity l11_14) have P4: "D C A CongA D C P" by (metis Col_def assms(2) assms(3) assms(4) bet_out_1 col124__nos l6_6 out2__conga out_trivial) have P5: "A B' Par C D" proof - have "\ Col B P A" using assms(4) col123__nos not_col_permutation_2 by blast then have "P A TS B B'" by (metis P2 assms(4) bet__ts cong_reverse_identity invert_two_sides not_col_permutation_3 os_distincts) then have "A C TS B' D" by (meson assms(2) assms(3) assms(4) bet_col bet_col1 col_preserves_two_sides l9_2 l9_8_2) moreover have "B' A C CongA D C A" proof - have "B' A C CongA B A P" by (simp add: P3 conga_comm conga_sym) moreover have "B A P CongA D C A" using P4 assms(5) not_conga not_conga_sym by blast ultimately show ?thesis using not_conga by blast qed ultimately show ?thesis using l12_21_b by blast qed have "C D Par A B" proof - have "A \ B" using assms(4) os_distincts by blast moreover have "C D Par A B'" using P5 par_symmetry by blast moreover have "Col A B' B" by (simp add: Col_def P2) ultimately show ?thesis using par_col_par by blast qed thus ?thesis using Par_cases by blast qed lemma l12_22_b: assumes "P Out A C" and "P A OS B D" and "B A P CongA D C P" shows "A B Par C D" proof cases assume "A = C" then show ?thesis using assms(2) assms(3) conga_comm conga_os__out not_par_not_col os_distincts out_col by blast next assume P1: "A \ C" { assume "Bet P A C" then have "A B Par C D" using P1 assms(2) assms(3) conga_diff2 l12_22_aux by blast } { assume P2: "Bet P C A" have "C D Par A B" proof - have "P C OS D B" using assms(1) assms(2) col_one_side one_side_symmetry out_col out_diff2 by blast moreover have "D C P CongA B A P" using assms(3) not_conga_sym by blast then show ?thesis by (metis P1 P2 assms(1) calculation l12_22_aux out_distinct) qed then have "A B Par C D" using Par_cases by auto } then show ?thesis using Out_def \Bet P A C \ A B Par C D\ assms(1) by blast qed lemma par_strict_par: assumes "A B ParStrict C D" shows "A B Par C D" using Par_def assms by auto lemma par_strict_distinct: assumes "A B ParStrict C D" shows " A \ B \ C \ D" using assms par_strict_neq1 par_strict_neq2 by auto lemma col_par: assumes "A \ B" and "B \ C" and "Col A B C" shows "A B Par B C" by (simp add: Par_def assms(1) assms(2) assms(3) col_trivial_1) lemma acute_col_perp__out: assumes "Acute A B C" and "Col B C A'" and "B C Perp A A'" shows "B Out A' C" proof - { assume P1: "\ Col B C A" then obtain B' where P2: "B C Perp B' B \ B C OS A B'" using assms(2) l10_15 os_distincts by blast have P3: "\ Col B' B C" using P2 col124__nos col_permutation_1 by blast { assume "Col B B' A" then have "A B C LtA A B C" using P2 acute_one_side_aux acute_sym assms(1) one_side_not_col124 by blast then have "False" by (simp add: nlta) } then have P4: "\ Col B B' A" by auto have P5: "B B' ParStrict A A'" proof - have "B B' Par A A'" proof - have "Coplanar B C B A" using ncop_distincts by blast moreover have "Coplanar B C B A'" using ncop_distincts by blast moreover have "Coplanar B C B' A" using P2 coplanar_perm_1 os__coplanar by blast moreover have "Coplanar B C B' A'" using assms(2) ncop__ncols by auto moreover have "B B' Perp B C" using P2 Perp_perm by blast moreover have "A A' Perp B C" using Perp_perm assms(3) by blast ultimately show ?thesis using l12_9 by auto qed moreover have "Col A A' A" by (simp add: col_trivial_3) moreover have "\ Col B B' A" by (simp add: P4) ultimately show ?thesis using par_not_col_strict by auto qed then have P6: "\ Col B B' A'" using P5 par_strict_not_col_4 by auto then have "B B' OS A' C" proof - have "B B' OS A' A" using P5 l12_6 one_side_symmetry by blast moreover have "B B' OS A C" using P2 acute_one_side_aux acute_sym assms(1) one_side_symmetry by blast ultimately show ?thesis using one_side_transitivity by blast qed then have "B Out A' C" using Col_cases assms(2) col_one_side_out by blast } then show ?thesis using assms(2) assms(3) perp_not_col2 by blast qed lemma acute_col_perp__out_1: assumes "Acute A B C" and "Col B C A'" and "B A Perp A A'" shows "B Out A' C" proof - obtain A0 where P1: "Bet A B A0 \ Cong B A0 A B" using segment_construction by blast obtain C0 where P2: "Bet C B C0 \ Cong B C0 C B" using segment_construction by blast have P3: "\ Col B A A'" using assms(3) col_trivial_2 perp_not_col2 by blast have "Bet A' B C0" proof - have P4: "Col A' B C0" using P2 acute_distincts assms(1) assms(2) bet_col col_transitivity_2 not_col_permutation_4 by blast { assume P5: "B Out A' C0" have "B Out A A0" proof - have "Bet C B A'" by (smt Bet_perm Col_def P2 P5 assms(2) between_exchange3 not_bet_and_out outer_transitivity_between2) then have "A B C CongA A0 B A'" using P1 P3 acute_distincts assms(1) cong_diff_4 l11_14 not_col_distincts by blast then have "Acute A' B A0" using acute_conga__acute acute_sym assms(1) by blast moreover have "B A0 Perp A' A" proof - have "B \ A0" using P1 P3 col_trivial_1 cong_reverse_identity by blast moreover have "B A Perp A' A" using Perp_perm assms(3) by blast moreover have "Col B A A0" using P1 bet_col not_col_permutation_4 by blast ultimately show ?thesis using perp_col by blast qed ultimately show ?thesis using Col_cases P1 acute_col_perp__out bet_col by blast qed then have "False" using P1 not_bet_and_out by blast } moreover then have "\ B Out A' C0" by auto ultimately show ?thesis using l6_4_2 P4 by blast qed then show ?thesis by (metis P2 P3 acute_distincts assms(1) cong_diff_3 l6_2 not_col_distincts) qed lemma conga_inangle_per2__inangle: assumes "Per A B C" and "T InAngle A B C" and "P B A CongA P B C" and "Per B P T" and "Coplanar A B C P" shows "P InAngle A B C" proof cases assume "P = T" then show ?thesis by (simp add: assms(2)) next assume P1: "P \ T" obtain P' where P2: "P' InAngle A B C \ P' B A CongA P' B C" using CongA_def angle_bisector assms(3) by presburger have P3: "Acute P' B A" using P2 acute_sym assms(1) conga_inangle_per__acute by blast have P4: "\ Col A B C" using assms(1) assms(3) conga_diff2 conga_diff56 l8_9 by blast have P5: "Col B P P'" proof - have "\ B Out A C" using Col_cases P4 out_col by blast moreover have "Coplanar A B P P'" proof - have T1: "\ Col C A B" using Col_perm P4 by blast moreover have "Coplanar C A B P" using assms(5) ncoplanar_perm_8 by blast moreover have "Coplanar C A B P'" using P2 inangle__coplanar ncoplanar_perm_21 by blast ultimately show ?thesis using coplanar_trans_1 by blast qed moreover have "Coplanar B C P P'" proof - have "Coplanar A B C P" by (meson P2 bet__coplanar calculation(1) calculation(2) col_in_angle_out coplanar_perm_18 coplanar_trans_1 inangle__coplanar l11_21_a l6_6 l6_7 not_col_permutation_4 not_col_permutation_5) have "Coplanar A B C P'" using P2 inangle__coplanar ncoplanar_perm_18 by blast then show ?thesis using P4 \Coplanar A B C P\ coplanar_trans_1 by blast qed ultimately show ?thesis using conga2_cop2__col P2 assms(3) by blast qed have "B Out P P'" proof - have "Acute T B P'" using P2 acute_sym assms(1) assms(2) conga_inangle2_per__acute by blast moreover have "B P' Perp T P" by (metis P1 P5 acute_distincts assms(3) assms(4) calculation col_per_perp conga_distinct l8_2 not_col_permutation_4) ultimately show ?thesis using Col_cases P5 acute_col_perp__out by blast qed then show ?thesis using Out_cases P2 in_angle_trans inangle_distincts out341__inangle by blast qed lemma perp_not_par: assumes "A B Perp X Y" shows "\ A B Par X Y" proof - obtain P where P1: "P PerpAt A B X Y" using Perp_def assms by blast { assume P2: "A B Par X Y" { assume P3: "A B ParStrict X Y" then have "False" proof - have "Col P A B" using Col_perm P1 perp_in_col by blast moreover have "Col P X Y" using P1 col_permutation_2 perp_in_col by blast ultimately show ?thesis using P3 par_not_col by blast qed } { assume P4: "A \ B \ X \ Y \ Col A X Y \ Col B X Y" then have "False" proof cases assume "A = Y" thus ?thesis using P4 assms not_col_permutation_1 perp_not_col by blast next assume "A \ Y" thus ?thesis using Col_perm P4 Perp_perm assms perp_not_col2 by blast qed } then have "False" using Par_def P2 \A B ParStrict X Y \ False\ by auto } thus ?thesis by auto qed lemma cong_conga_perp: assumes "B P TS A C" and "Cong A B C B" and "A B P CongA C B P" shows "A C Perp B P" proof - have P1: " \ Col A B P" using TS_def assms(1) by blast then have P2: "B \ P" using col_trivial_2 by blast have P3: "A \ B" using assms(1) ts_distincts by blast have P4: "C \ B" using assms(1) ts_distincts by auto have P5: "A \ C" using assms(1) not_two_sides_id by auto show ?thesis proof cases assume P6: "Bet A B C" then have "Per P B A" by (meson Tarski_neutral_dimensionless.conga_comm Tarski_neutral_dimensionless_axioms assms(3) l11_18_2) then show ?thesis using P2 P3 P5 Per_perm P6 bet_col per_perp perp_col by blast next assume P7: "\ Bet A B C" obtain T where P7A: "Col T B P \ Bet A T C" using TS_def assms(1) by auto then have P8: "B \ T" using P7 by blast then have P9: "T B A CongA T B C" by (meson Col_cases P7A Tarski_neutral_dimensionless.col_conga__conga Tarski_neutral_dimensionless.conga_comm Tarski_neutral_dimensionless_axioms assms(3)) then have P10: "Cong T A T C" using assms(2) cong2_conga_cong cong_reflexivity not_cong_2143 by blast then have P11: "T Midpoint A C" using P7A midpoint_def not_cong_2134 by blast have P12: "Per B T A" using P11 Per_def assms(2) not_cong_2143 by blast then show ?thesis proof - have "A C Perp B T" by (metis P11 P12 P5 P8 col_per_perp midpoint_col midpoint_distinct_1) moreover have "B \ T" by (simp add: P8) moreover have "T \ A" using P1 P7A by blast moreover have "C \ T" using P10 P5 cong_identity by blast moreover have "C \ A" using P5 by auto moreover have "Col T A C" by (meson P7A bet_col not_col_permutation_4) ultimately show ?thesis using P2 P7A not_col_permutation_4 perp_col1 by blast qed qed qed lemma perp_inter_exists: assumes "A B Perp C D" shows "\ P. Col A B P \ Col C D P" proof - obtain P where "P PerpAt A B C D" using Perp_def assms by auto then show ?thesis using perp_in_col by blast qed lemma perp_inter_perp_in: assumes "A B Perp C D" shows "\ P. Col A B P \ Col C D P \ P PerpAt A B C D" by (meson Perp_def Tarski_neutral_dimensionless.perp_in_col Tarski_neutral_dimensionless_axioms assms) end context Tarski_2D begin lemma l12_9_2D: assumes "A1 A2 Perp C1 C2" and "B1 B2 Perp C1 C2" shows "A1 A2 Par B1 B2" using l12_9 all_coplanar assms(1) assms(2) by auto end context Tarski_neutral_dimensionless begin subsection "Tarski: Chapter 13" subsubsection "Introduction" lemma per2_col_eq: assumes "A \ P" and "A \ P'" and "Per A P B" and "Per A P' B" and "Col P A P'" shows "P = P'" by (metis assms(1) assms(2) assms(3) assms(4) assms(5) col_per2_cases l6_16_1 l8_2 not_col_permutation_3) lemma per2_preserves_diff: assumes "PO \ A'" and "PO \ B'" and "Col PO A' B'" and "Per PO A' A" and "Per PO B' B" and "A' \ B'" shows "A \ B" using assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) not_col_permutation_4 per2_col_eq by blast lemma per23_preserves_bet: assumes "Bet A B C" and "A \ B'" and "A \ C'" and "Col A B' C'" and "Per A B' B" and "Per A C' C" shows "Bet A B' C'" proof - have P1: "Col A B C" by (simp add: assms(1) bet_col) show ?thesis proof cases assume P2: "B = B'" then have "Col A C' C" using P1 assms(2) assms(4) col_transitivity_1 by blast then have P4: "A = C' \ C = C'" by (simp add: assms(6) l8_9) { assume "A = C'" then have "Bet A B' C'" using assms(3) by auto } { assume "C = C'" then have "Bet A B' C'" using P2 assms(1) by auto } then show ?thesis using P4 assms(3) by auto next assume T1: "B \ B'" have T2: "A \ C" using assms(3) assms(6) l8_8 by auto have T3: "C \ C'" using P1 T1 assms(2) assms(3) assms(4) assms(5) col_trivial_3 colx l8_9 not_col_permutation_5 by blast have T3A: "A B' Perp B' B" using T1 assms(2) assms(5) per_perp by auto have T3B: "A C' Perp C' C" using T3 assms(3) assms(6) per_perp by auto have T4: "B B' Par C C'" proof - have "Coplanar A B' B C" using P1 ncop__ncols by blast moreover have "Coplanar A B' B C'" using assms(4) ncop__ncols by blast moreover have "Coplanar A B' B' C" using ncop_distincts by blast moreover have "B B' Perp A B'" using Perp_perm \A B' Perp B' B\ by blast moreover have "C C' Perp A B'" using Col_cases Perp_cases T3B assms(2) assms(4) perp_col1 by blast ultimately show ?thesis using l12_9 bet__coplanar between_trivial by auto qed moreover have "Bet A B' C'" proof cases assume "B = C" then show ?thesis by (metis T1 Tarski_neutral_dimensionless.per_col_eq Tarski_neutral_dimensionless_axioms assms(4) assms(5) calculation l6_16_1 l6_6 or_bet_out out_diff1 par_id) next assume T6: "B \ C" have T7: "\ Col A B' B" using T1 assms(2) assms(5) l8_9 by blast have T8: "\ Col A C' C" using T3 assms(3) assms(6) l8_9 by blast have T9: "B' \ C'" using P1 T6 assms(2) assms(5) assms(6) col_per2__per col_permutation_1 l8_2 l8_8 by blast have T10: "B B' ParStrict C C' \ (B \ B' \ C \ C' \ Col B C C' \ Col B' C C')" using Par_def calculation by blast { assume T11: "B B' ParStrict C C'" then have T12: "B B' OS C' C" using l12_6 one_side_symmetry by blast have "B B' TS A C" using Col_cases T6 T7 assms(1) bet__ts by blast then have "Bet A B' C'" using T12 assms(4) l9_5 l9_9 not_col_distincts or_bet_out by blast } { assume "B \ B' \ C \ C' \ Col B C C' \ Col B' C C'" then have "Bet A B' C'" using Col_def T6 T8 assms(1) col_transitivity_2 by blast } then show ?thesis using T10 \B B' ParStrict C C' \ Bet A B' C'\ by blast qed ultimately show ?thesis by (smt P1 Par_def T1 T2 assms(4) col_transitivity_2 not_col_permutation_1 par_strict_not_col_2) qed qed lemma per23_preserves_bet_inv: assumes "Bet A B' C'" and "A \ B'" and "Col A B C" and "Per A B' B" and "Per A C' C" shows "Bet A B C" proof cases assume T1: "B = B'" then have "Col A C' C" using Col_def assms(1) assms(2) assms(3) col_transitivity_1 by blast then have T2: "A = C' \ C = C'" by (simp add: assms(5) l8_9) { assume "A = C'" then have "Bet A B C" using assms(1) assms(2) between_identity by blast } { assume "C = C'" then have "Bet A B C" by (simp add: T1 assms(1)) } then show ?thesis using T2 \A = C' \ Bet A B C\ by auto next assume P1: "B \ B'" then have P2: "A B' Perp B' B" using assms(2) assms(4) per_perp by auto have "Per A C' C" by (simp add: assms(5)) then have P2: "C' PerpAt A C' C' C" by (metis (mono_tags, lifting) Col_cases P1 assms(1) assms(2) assms(3) assms(4) bet_col bet_neq12__neq col_transitivity_1 l8_9 per_perp_in) then have P3: "A C' Perp C' C" using perp_in_perp by auto then have "C' \ C" using \A C' Perp C' C\ perp_not_eq_2 by auto have "C' PerpAt C' A C C'" by (simp add: Perp_in_perm P2) then have "(C' A Perp C C') \ (C' C' Perp C C')" using Perp_def by blast have "A \ C'" using assms(1) assms(2) between_identity by blast { assume "C' A Perp C C'" have "Col A B' C'" using assms(1) by (simp add: Col_def) have "A B' Perp C' C" using Col_cases \A C' Perp C' C\ \Col A B' C'\ assms(2) perp_col by blast have P7: "B' B Par C' C" proof - have "Coplanar A B' B' C'" using ncop_distincts by blast moreover have "Coplanar A B' B' C" using ncop_distincts by auto moreover have "Coplanar A B' B C'" using Bet_perm assms(1) bet__coplanar ncoplanar_perm_20 by blast moreover have "Coplanar A B' B C" using assms(3) ncop__ncols by auto moreover have "B' B Perp A B'" by (metis P1 Perp_perm assms(2) assms(4) per_perp) moreover have "C' C Perp A B'" using Perp_cases \A B' Perp C' C\ by auto ultimately show ?thesis using l12_9 by blast qed have "Bet A B C" proof cases assume "B = C" then show ?thesis by (simp add: between_trivial) next assume T1: "B \ C" have T2: "B' B ParStrict C' C \ (B' \ B \ C' \ C \ Col B' C' C \ Col B C' C)" using P7 Par_def by auto { assume T3: "B' B ParStrict C' C" then have "B' \ C'" using not_par_strict_id by auto have "\ X. Col X B' B \ Col X B' C" using col_trivial_1 by blast have "B' B OS C' C" by (simp add: T3 l12_6) have "B' B TS A C'" by (metis Bet_cases T3 assms(1) assms(2) bet__ts l9_2 par_strict_not_col_1) then have T8: "B' B TS C A" using \B' B OS C' C\ l9_2 l9_8_2 by blast then obtain T where T9: "Col T B' B \ Bet C T A" using TS_def by auto have "\ Col A C B'" using T8 assms(3) not_col_permutation_2 not_col_permutation_3 ts__ncol by blast then have "T = B" by (metis Col_def Col_perm T9 assms(3) colx) then have "Bet A B C" using Bet_cases T9 by auto } { assume "B' \ B \ C' \ C \ Col B' C' C \ Col B C' C" then have "Col A B' B" by (metis Col_perm T1 assms(3) l6_16_1) then have "A = B' \ B = B'" using assms(4) l8_9 by auto then have "Bet A B C" by (simp add: P1 assms(2)) } then show ?thesis using T2 \B' B ParStrict C' C \ Bet A B C\ by auto qed } then show ?thesis by (simp add: P3 perp_comm) qed lemma per13_preserves_bet: assumes "Bet A B C" and "B \ A'" and "B \ C'" and "Col A' B C'" and "Per B A' A" and "Per B C' C" shows "Bet A' B C'" by (smt Col_cases Tarski_neutral_dimensionless.per23_preserves_bet_inv Tarski_neutral_dimensionless_axioms assms(1) assms(4) assms(5) assms(6) bet_col between_equality between_symmetry per_distinct third_point) lemma per13_preserves_bet_inv: assumes "Bet A' B C'" and "B \ A'" and "B \ C'" and "Col A B C" and "Per B A' A" and "Per B C' C" shows "Bet A B C" proof - have P1: "Col A' B C'" by (simp add: Col_def assms(1)) show ?thesis proof cases assume "A = A'" then show ?thesis using P1 assms(1) assms(3) assms(4) assms(6) col_transitivity_2 l8_9 not_bet_distincts by blast next assume "A \ A'" show ?thesis by (metis Col_cases P1 Tarski_neutral_dimensionless.per23_preserves_bet Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) between_equality between_symmetry third_point) qed qed lemma per3_preserves_bet1: assumes "Col PO A B" and "Bet A B C" and "PO \ A'" and "PO \ B'" and "PO \ C'" and "Per PO A' A" and "Per PO B' B" and "Per PO C' C" and "Col A' B' C'" and "Col PO A' B'" shows "Bet A' B' C'" proof cases assume "A = B" then show ?thesis using assms(10) assms(3) assms(4) assms(6) assms(7) between_trivial2 per2_preserves_diff by blast next assume P1: "A \ B" show ?thesis proof cases assume P2: "A = A'" show ?thesis proof cases assume P3: "B = B'" then have "Col PO C C'" by (metis (no_types, opaque_lifting) Col_def P1 P2 assms(1) assms(2) assms(9) col_transitivity_1) then have "C = C'" using assms(5) assms(8) l8_9 not_col_permutation_5 by blast then show ?thesis using P2 P3 assms(2) by blast next assume P4: "B \ B'" show ?thesis proof cases assume "A = B'" then show ?thesis using P2 between_trivial2 by auto next assume "A \ B'" have "A \ C" using P1 assms(2) between_identity by blast have P7: "\ Col PO B' B" using P4 assms(4) assms(7) l8_9 by blast show ?thesis using P2 P7 assms(1) assms(10) assms(3) col_transitivity_1 by blast qed qed next assume R1: "A \ A'" show ?thesis proof cases assume R2: "A' = B'" then show ?thesis by (simp add: between_trivial2) next assume R3: "A' \ B'" show ?thesis proof cases assume "B = C" have "B' = C'" by (metis Tarski_neutral_dimensionless.per2_col_eq Tarski_neutral_dimensionless_axioms \A' \ B'\ \B = C\ assms(10) assms(4) assms(5) assms(7) assms(8) assms(9) col_transitivity_2 not_col_permutation_2) then show ?thesis by (simp add: between_trivial) next assume R4: "B \ C" show ?thesis proof cases assume "B = B'" then show ?thesis by (metis R1 assms(1) assms(10) assms(3) assms(4) assms(6) l6_16_1 l8_9 not_col_permutation_2) next assume R5: "B \ B'" show ?thesis proof cases assume "A' = B" then show ?thesis using R5 assms(10) assms(4) assms(7) col_permutation_5 l8_9 by blast next assume R5A: "A' \ B" have R6: "C \ C'" by (metis P1 R1 R3 assms(1) assms(10) assms(2) assms(3) assms(5) assms(6) assms(9) bet_col col_permutation_1 col_trivial_2 l6_21 l8_9) have R7: "A A' Perp PO A'" by (metis Perp_cases R1 assms(3) assms(6) per_perp) have R8: "C C' Perp PO A'" by (smt Perp_cases R3 R6 assms(10) assms(3) assms(5) assms(8) assms(9) col2__eq col3 col_per_perp col_trivial_2 l8_2 per_perp) have "A A' Par C C'" proof - have "Coplanar PO A' A C" using P1 assms(1) assms(2) bet_col col_trivial_2 colx ncop__ncols by blast moreover have "Coplanar PO A' A C'" using R3 assms(10) assms(9) col_trivial_2 colx ncop__ncols by blast moreover have "Coplanar PO A' A' C" using ncop_distincts by blast moreover have "Coplanar PO A' A' C'" using ncop_distincts by blast ultimately show ?thesis using l12_9 R7 R8 by blast qed have S1: "B B' Perp PO A'" by (metis Col_cases Per_cases Perp_perm R5 assms(10) assms(3) assms(4) assms(7) col_per_perp) have "A A' Par B B'" proof - have "Coplanar PO A' A B" using assms(1) ncop__ncols by auto moreover have "Coplanar PO A' A B'" using assms(10) ncop__ncols by auto moreover have "Coplanar PO A' A' B" using ncop_distincts by auto moreover have "Coplanar PO A' A' B'" using ncop_distincts by auto moreover have "A A' Perp PO A'" by (simp add: R7) moreover have "B B' Perp PO A'" by (simp add: S1) ultimately show ?thesis using l12_9 by blast qed { assume "A A' ParStrict B B'" then have "A A' OS B B'" by (simp add: l12_6) have "B B' TS A C" using R4 \A A' ParStrict B B'\ assms(2) bet__ts par_strict_not_col_3 by auto have "B B' OS A A'" using \A A' ParStrict B B'\ pars__os3412 by auto have "B B' TS A' C" using \B B' OS A A'\ \B B' TS A C\ l9_8_2 by blast have "Bet A' B' C'" proof cases assume "C = C'" then show ?thesis using R6 by auto next assume "C \ C'" have "C C' Perp PO A'" by (simp add: R8) have Q2: "B B' Par C C'" proof - have "Coplanar PO A' B C" by (metis P1 assms(1) assms(2) bet_col col_transitivity_1 colx ncop__ncols not_col_permutation_5) moreover have "Coplanar PO A' B C'" using R3 assms(10) assms(9) col_trivial_2 colx ncop__ncols by blast moreover have "Coplanar PO A' B' C" by (simp add: assms(10) col__coplanar) moreover have "Coplanar PO A' B' C'" using assms(10) col__coplanar by auto moreover have "B B' Perp PO A'" by (simp add: S1) moreover have "C C' Perp PO A'" by (simp add: R8) ultimately show ?thesis using l12_9 by auto qed then have Q3: "(B B' ParStrict C C') \ (B \ B' \ C \ C' \ Col B C C' \ Col B' C C')" by (simp add: Par_def) { assume "B B' ParStrict C C'" then have "B B' OS C C'" using l12_6 by auto then have "B B' TS C' A'" using \B B' TS A' C\ l9_2 l9_8_2 by blast then obtain T where Q4: "Col T B B' \ Bet C' T A'" using TS_def by blast have "T = B'" proof - have "\ Col B B' A'" using \B B' OS A A'\ col124__nos by auto moreover have "A' \ C'" using \B B' TS C' A'\ not_two_sides_id by auto moreover have "Col B B' T" using Col_cases Q4 by auto moreover have "Col B B' B'" using not_col_distincts by blast moreover have "Col A' C' T" by (simp add: Col_def Q4) ultimately show ?thesis by (meson assms(9) col_permutation_5 l6_21) qed then have "Bet A' B' C'" using Q4 between_symmetry by blast } { assume "B \ B' \ C \ C' \ Col B C C' \ Col B' C C'" then have "Bet A' B' C'" using TS_def \B B' TS A C\ l6_16_1 not_col_permutation_2 by blast } then show ?thesis using Q3 \B B' ParStrict C C' \ Bet A' B' C'\ by blast qed } { assume R8: "A \ A' \ B \ B' \ Col A B B' \ Col A' B B'" have "A' A Perp PO A'" by (simp add: R7 perp_left_comm) have "\ Col A' A PO" using Col_cases R8 assms(3) assms(6) l8_9 by blast then have "Bet A' B' C'" using Col_perm P1 R8 assms(1) l6_16_1 by blast } then show ?thesis using Par_def \A A' Par B B'\ \A A' ParStrict B B' \ Bet A' B' C'\ by auto qed qed qed qed qed qed lemma per3_preserves_bet2_aux: assumes "Col PO A C" and "A \ C'" and "Bet A B' C'" and "PO \ A" and "PO \ B'" and "PO \ C'" and "Per PO B' B" and "Per PO C' C" and "Col A B C" and "Col PO A C'" shows "Bet A B C" proof cases assume "A = B" then show ?thesis by (simp add: between_trivial2) next assume P1: "A \ B" show ?thesis proof cases assume "B = C" then show ?thesis by (simp add: between_trivial) next assume P2: "B \ C" have P3: "Col PO A B'" by (metis Col_def assms(10) assms(2) assms(3) l6_16_1) then have P4: "Col PO B' C'" using assms(10) assms(4) col_transitivity_1 by blast show ?thesis proof cases assume "B = B'" thus ?thesis by (metis Tarski_neutral_dimensionless.per_col_eq Tarski_neutral_dimensionless_axioms assms(1) assms(10) assms(3) assms(4) assms(6) assms(8) col_transitivity_1) next assume P5: "B \ B'" have P6: "C = C'" using assms(1) assms(10) assms(4) assms(6) assms(8) col_transitivity_1 l8_9 by blast then have "False" by (metis P3 P5 P6 Tarski_neutral_dimensionless.per_col_eq Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(4) assms(5) assms(7) assms(9) col_transitivity_1 l6_16_1 not_col_permutation_4) then show ?thesis by blast qed qed qed lemma per3_preserves_bet2: assumes "Col PO A C" and "A' \ C'" and "Bet A' B' C'" and "PO \ A'" and "PO \ B'" and "PO \ C'" and "Per PO A' A" and "Per PO B' B" and "Per PO C' C" and "Col A B C" and "Col PO A' C'" shows "Bet A B C" proof cases assume "A = A'" then show ?thesis using assms(1) assms(10) assms(11) assms(2) assms(3) assms(4) assms(5) assms(6) assms(8) assms(9) per3_preserves_bet2_aux by blast next assume P1: "A \ A'" show ?thesis proof cases assume "C = C'" thus ?thesis by (metis P1 assms(1) assms(11) assms(4) assms(6) assms(7) col_trivial_3 l6_21 l8_9 not_col_permutation_2) next assume P2: "C \ C'" then have P3: "PO A' Perp C C'" by (metis assms(11) assms(4) assms(6) assms(9) col_per_perp l8_2 not_col_permutation_1) have P4: "PO A' Perp A A'" using P1 assms(4) assms(7) per_perp perp_right_comm by auto have "A A' Par C C'" proof - have "Coplanar PO A' A C" using assms(1) ncop__ncols by blast moreover have "Coplanar PO A' A C'" by (meson assms(11) ncop__ncols) moreover have "Coplanar PO A' A' C" using ncop_distincts by blast moreover have "Coplanar PO A' A' C'" using ncop_distincts by blast moreover have "A A' Perp PO A'" using P4 Perp_cases by blast moreover have "C C' Perp PO A'" using P3 Perp_cases by auto ultimately show ?thesis using l12_9 by blast qed { assume P5: "A A' ParStrict C C'" then have P6: "A A' OS C C'" by (simp add: l12_6) have P7: "C C' OS A A'" by (simp add: P5 pars__os3412) have "Bet A B C" proof cases assume P8: "B = B'" then have "A' A OS B C'" by (metis P6 assms(10) assms(3) bet_out col123__nos col124__nos invert_one_side out_one_side) then have "A A' OS B C'" by (simp add: invert_one_side) then have "A A' OS B C" using P6 one_side_symmetry one_side_transitivity by blast then have P12: "A Out B C" using assms(10) col_one_side_out by blast have "C' C OS B A'" by (metis Col_perm P5 P7 P8 assms(10) assms(3) bet_out_1 col123__nos out_one_side par_strict_not_col_2) then have "C C' OS B A" by (meson P7 invert_one_side one_side_symmetry one_side_transitivity) then have "C C' OS A B" using one_side_symmetry by blast then have "C Out A B" using assms(10) col_one_side_out col_permutation_2 by blast then show ?thesis by (simp add: P12 out2__bet) next assume T1: "B \ B'" have T2: "PO A' Perp B B'" proof - have "Per PO B' B" by (simp add: assms(8)) then have "B' PerpAt PO B' B' B" using T1 assms(5) per_perp_in by auto then have "B' PerpAt B' PO B B'" by (simp add: perp_in_comm) then have T4: "B' PO Perp B B' \ B' B' Perp B B'" using Perp_def by auto { assume T5: "B' PO Perp B B'" have "Col A' B' C'" by (simp add: assms(3) bet_col) then have "Col PO B' A'" using assms(11) assms(2) col2__eq col_permutation_4 col_permutation_5 by blast then have "PO A' Perp B B'" by (metis T5 assms(4) col_trivial_3 perp_col2 perp_comm) } { assume "B' B' Perp B B'" then have "PO A' Perp B B'" using perp_distinct by auto } then show ?thesis using T4 \B' PO Perp B B' \ PO A' Perp B B'\ by linarith qed have T6: "B B' Par A A'" proof - have "Coplanar PO A' B A" by (metis Col_cases P7 assms(1) assms(10) col_transitivity_2 ncop__ncols os_distincts) moreover have "Coplanar PO A' B A'" using ncop_distincts by blast moreover have "Coplanar PO A' B' A" proof - have "(Bet PO A' C' \ Bet PO C' A') \ Bet C' PO A'" by (meson assms(11) third_point) then show ?thesis by (meson Bet_perm assms(3) bet__coplanar between_exchange2 l5_3 ncoplanar_perm_8) qed moreover have "Coplanar PO A' B' A'" using ncop_distincts by auto moreover have "B B' Perp PO A'" using Perp_cases T2 by blast moreover have "A A' Perp PO A'" using P4 Perp_cases by blast ultimately show ?thesis using l12_9 by blast qed { assume "B B' ParStrict A A'" then have "B B' OS A A'" by (simp add: l12_6) have "B B' Par C C'" proof - have "Coplanar PO A' B C" by (metis Col_cases P7 assms(1) assms(10) col2__eq ncop__ncols os_distincts) moreover have "Coplanar PO A' B C'" using assms(11) ncop__ncols by auto moreover have "Coplanar PO A' B' C" by (metis Out_def assms(11) assms(2) assms(3) col_trivial_2 l6_16_1 ncop__ncols not_col_permutation_1 out_col) moreover have "Coplanar PO A' B' C'" using assms(11) ncop__ncols by blast moreover have "B B' Perp PO A'" using Perp_cases T2 by blast moreover have "C C' Perp PO A'" using P3 Perp_cases by auto ultimately show ?thesis using l12_9 by blast qed { assume T9: "B B' ParStrict C C'" then have T10: "B B' OS C C'" by (simp add: l12_6) have T11: "B B' TS A' C'" by (metis Col_cases T10 \B B' ParStrict A A'\ assms(3) bet__ts invert_two_sides os_distincts par_strict_not_col_4) have T12: "B B' TS A C'" using \B B' OS A A'\ \B B' TS A' C'\ l9_8_2 one_side_symmetry by blast then have T12A: "B B' TS C A" using T10 l9_2 l9_8_2 one_side_symmetry by blast then obtain T where T13: "Col T B B' \ Bet C T A" using TS_def by auto then have "B = T" by (metis Col_perm TS_def T12A assms(10) bet_col1 col_transitivity_2 col_two_sides_bet) then have "Bet A B C" using Bet_perm T13 by blast } { assume "B \ B' \ C \ C' \ Col B C C' \ Col B' C C'" then have "Bet A B C" by (metis Col_cases P5 assms(10) col3 col_trivial_2 not_bet_distincts par_strict_not_col_3) } then have "Bet A B C" using Par_def \B B' Par C C'\ \B B' ParStrict C C' \ Bet A B C\ by auto } { assume "B \ B' \ A \ A' \ Col B A A' \ Col B' A A'" then have "Bet A B C" by (smt P6 assms(10) col123__nos l6_16_1 not_bet_distincts not_col_permutation_1) } then show ?thesis using Par_def T6 \B B' ParStrict A A' \ Bet A B C\ by auto qed } { assume "A \ A' \ C \ C' \ Col A C C' \ Col A' C C'" then have "Bet A B C" by (metis Col_perm P3 Par_def assms(11) assms(2) assms(4) col_transitivity_1 perp_not_par) } thus ?thesis using Par_def \A A' Par C C'\ \A A' ParStrict C C' \ Bet A B C\ by auto qed qed lemma symmetry_preserves_per: assumes "Per B P A" and "B Midpoint A A'" and "B Midpoint P P'" shows "Per B P' A'" proof - obtain C where P1: "P Midpoint A C" using symmetric_point_construction by blast obtain C' where P2: "B Midpoint C C'" using symmetric_point_construction by blast have P3: "P' Midpoint A' C'" using P1 P2 assms(2) assms(3) symmetry_preserves_midpoint by blast have "Cong B A' B C'" by (meson P1 P2 assms(1) assms(2) l7_16 l7_3_2 per_double_cong) then show ?thesis using P3 Per_def by blast qed lemma l13_1_aux: assumes "\ Col A B C" and "P Midpoint B C" and "Q Midpoint A C" and "R Midpoint A B" shows "\ X Y. (R PerpAt X Y A B \ X Y Perp P Q \ Coplanar A B C X \ Coplanar A B C Y)" proof - have P1: "Q \ C" using assms(1) assms(3) midpoint_not_midpoint not_col_distincts by blast have P2: "P \ C" using assms(1) assms(2) is_midpoint_id_2 not_col_distincts by blast then have "Q \ R" using assms(2) assms(3) assms(4) l7_3 symmetric_point_uniqueness by blast have "R \ B" using assms(1) assms(4) midpoint_not_midpoint not_col_distincts by blast { assume V1: "Col P Q C" have V2: "Col B P C" by (simp add: assms(2) bet_col midpoint_bet) have V3: "Col A Q C" by (simp add: assms(3) bet_col midpoint_bet) have "Col A R B" using assms(4) midpoint_col not_col_permutation_4 by blast then have "Col A B C" using V1 V2 V3 by (metis P1 P2 col2__eq col_permutation_5) then have "False" using assms(1) by auto } then have P2A: "\ Col P Q C" by auto then obtain C' where P3: "Col P Q C' \ P Q Perp C C'" using l8_18_existence by blast obtain A' where P4: "Q Midpoint C' A'" using symmetric_point_construction by auto obtain B' where P5: "P Midpoint C' B'" using symmetric_point_construction by auto have P6: "Cong C C' B B'" using Mid_cases P5 assms(2) l7_13 by blast have P7: "Cong C C' A A'" using P4 assms(3) l7_13 l7_2 by blast have P8: "Per P B' B" proof cases assume "P = C'" then show ?thesis using P5 Per_cases is_midpoint_id l8_5 by blast next assume "P \ C'" then have "P C' Perp C C'" using P3 perp_col by blast then have "Per P C' C" using Perp_perm perp_per_2 by blast then show ?thesis using symmetry_preserves_per Mid_perm P5 assms(2) by blast qed have P8A: "Per Q A' A" proof cases assume "Q = C'" then show ?thesis using P4 Per_cases is_midpoint_id l8_5 by blast next assume "Q \ C'" then have "C' Q Perp C C'" using P3 col_trivial_2 perp_col2 by auto then have "Per Q C' C" by (simp add: perp_per_1) then show ?thesis by (meson Mid_cases P4 assms(3) l7_3_2 midpoint_preserves_per) qed have P9: "Col A' C' Q" using P4 midpoint_col not_col_permutation_3 by blast have P10: "Col B' C' P" using P5 midpoint_col not_col_permutation_3 by blast have P11: "P \ Q" using P2A col_trivial_1 by auto then have P12: "A' \ B'" using P4 P5 l7_17 by blast have P13: "Col A' B' P" by (metis P10 P3 P4 P5 P9 col2__eq col_permutation_5 midpoint_distinct_1 not_col_distincts) have P14: "Col A' B' Q" by (smt P10 P3 P4 P5 P9 col3 col_permutation_1 midpoint_distinct_1 not_col_distincts) have P15: "Col A' B' C'" using P11 P13 P14 P3 colx by blast have P16: "C \ C'" using P2A P3 by blast then have P17: "A \ A'" using P7 cong_diff by blast have P18: "B \ B'" using P16 P6 cong_diff by blast have P19: "Per P A' A" proof cases assume P20: "A' = Q" then have "A' P Perp C A'" by (metis P3 P4 Perp_cases midpoint_not_midpoint) then have "Per P A' C" by (simp add: perp_per_1) then show ?thesis using P20 assms(3) l7_2 l8_4 by blast next assume "A' \ Q" then show ?thesis by (meson P12 P13 P14 P8A col_transitivity_1 l8_2 per_col) qed have "Per Q B' B" proof cases assume P21: "P = B'" then have P22: "C' = B'" using P5 is_midpoint_id_2 by auto then have "Per Q B' C" using P3 P21 perp_per_1 by auto thus ?thesis by (metis Col_perm P16 P21 P22 assms(2) midpoint_col per_col) next assume P23: "P \ B'" have "Col B' P Q" using P12 P13 P14 col_transitivity_2 by blast then have "Per B B' Q" using P8 P23 l8_2 l8_3 by blast thus ?thesis using Per_perm by blast qed then have P24: "Per A' B' B" using P11 P13 P14 P8 l8_3 not_col_permutation_2 by blast have P25: "Per A A' B'" using P11 P13 P14 P19 P8A l8_2 l8_3 not_col_permutation_5 by blast then have "Per B' A' A" using Per_perm by blast then have "\ Col B' A' A" using P12 P17 P25 per_not_col by auto then have P26: "\ Col A' B' A" using Col_cases by auto have "\ Col A' B' B" using P12 P18 P24 l8_9 by auto obtain X where P28: "X Midpoint A' B'" using midpoint_existence by blast then have P28A: "Col A' B' X" using midpoint_col not_col_permutation_2 by blast then have "\ Q. A' B' Perp Q X \ A' B' OS A Q" by (simp add: P26 l10_15) then obtain y where P29: "A' B' Perp y X \ A' B' OS A y" by blast then obtain B'' where P30: "(X y Perp A B'' \ A = B'') \ (\ M. (Col X y M \ M Midpoint A B''))" using ex_sym by blast then have P31: "B'' A ReflectL X y" using P30 ReflectL_def by blast have P32: "X \ y" using P29 P28A col124__nos by blast then have "X \ y \ B'' A ReflectL X y \ X = y \ X Midpoint A B''" using P31 by auto then have P33: "B'' A Reflect X y" by (simp add: Reflect_def) have P33A: "X \ y \ A' B' ReflectL X y" using P28 P29 Perp_cases ReflectL_def P32 col_trivial_3 l10_4_spec by blast then have P34: "A' B' Reflect X y" using Reflect_def by blast have P34A: "A B'' Reflect X y" using P33 l10_4 by blast then have P35: "Cong B'' B' A A'" using P34 l10_10 by auto have "Per A' B' B''" proof - have R1: "X \ y \ A B'' ReflectL X y \ X = y \ X Midpoint B'' A" by (simp add: P31 P32 l10_4_spec) have R2: "X \ y \ A' B' ReflectL X y \ X = y \ X Midpoint B' A'" using P33A by linarith { assume "X \ y \ A B'' ReflectL X y \ X \ y \ A' B' ReflectL X y" then have "Per A' B' B''" using \Per B' A' A\ image_spec_preserves_per l10_4_spec by blast } { assume "X \ y \ A B'' ReflectL X y \ X = y \ X Midpoint B' A'" then have "Per A' B' B''" by blast } { assume "X = y \ X Midpoint B'' A \ X \ y \ A' B' ReflectL X y" then have "Per A' B' B''" by blast } { assume "X = y \ X Midpoint B'' A \ X = y \ X Midpoint B' A'" then have "Per A' B' B''" using P32 by blast } then show ?thesis using R1 R2 using \X \ y \ A B'' ReflectL X y \ X \ y \ A' B' ReflectL X y \ Per A' B' B''\ by auto qed have "A' B' OS A B''" proof - { assume S1: "X y Perp A B''" have "Coplanar A y A' X" by (metis P28A P29 col_one_side coplanar_perm_16 ncop_distincts os__coplanar) have "Coplanar A y B' X" by (smt P12 P28A P29 col2_cop__cop col_transitivity_1 ncoplanar_perm_22 not_col_permutation_5 os__coplanar) have S2: "\ Col A X y" using Col_perm P34A S1 local.image_id perp_distinct by blast have "A' B' Par A B''" proof - have "Coplanar X y A' A" using \Coplanar A y A' X\ ncoplanar_perm_21 by blast moreover have "Coplanar X y A' B''" proof - have "Coplanar A X y A'" using \Coplanar X y A' A\ ncoplanar_perm_9 by blast moreover have "Coplanar A X y B''" using Coplanar_def S1 perp_inter_exists by blast ultimately show ?thesis using S2 coplanar_trans_1 by auto qed moreover have "Coplanar X y B' A" proof - have "\ Col A X y" by (simp add: S2) moreover have "Coplanar A X y B'" using \Coplanar A y B' X\ ncoplanar_perm_3 by blast moreover have "Coplanar A X y B''" using Coplanar_def S1 perp_inter_exists by blast ultimately show ?thesis using ncoplanar_perm_18 by blast qed moreover have "Coplanar X y B' B''" proof - have "\ Col A X y" by (simp add: S2) moreover have "Coplanar A X y B'" using \Coplanar X y B' A\ ncoplanar_perm_9 by blast moreover have "Coplanar A X y B''" using Coplanar_def S1 perp_inter_exists by blast ultimately show ?thesis using coplanar_trans_1 by blast qed ultimately show ?thesis using l12_9 using P29 Perp_cases S1 by blast qed have "A' B' OS A B''" proof - { assume "A' B' ParStrict A B''" have "A' B' OS A B''" using l12_6 using \A' B' ParStrict A B''\ by blast } { assume "A' \ B' \ A \ B'' \ Col A' A B'' \ Col B' A B''" have "A' B' OS A B''" using P26 \A' B' Par A B''\ \A' B' ParStrict A B'' \ A' B' OS A B''\ col_trivial_3 par_not_col_strict by blast } then show ?thesis using Par_def \A' B' Par A B''\ \A' B' ParStrict A B'' \ A' B' OS A B''\ by auto qed } { assume "A = B''" then have "A' B' OS A B''" using P12 P25 \Per A' B' B''\ l8_2 l8_7 by blast } then show ?thesis using P30 \X y Perp A B'' \ A' B' OS A B''\ by blast qed have "A' B' OS A B" proof - have "A' B' TS A C" proof - have "\ Col A A' B'" using Col_perm \\ Col B' A' A\ by blast moreover have "\ Col C A' B'" by (metis P13 P14 P2A \\ Col B' A' A\ col3 not_col_distincts not_col_permutation_3 not_col_permutation_4) moreover have "\ T. Col T A' B' \ Bet A T C" using P14 assms(3) midpoint_bet not_col_permutation_1 by blast ultimately show ?thesis by (simp add: TS_def) qed moreover have "A' B' TS B C" by (metis Col_cases P13 TS_def \\ Col A' B' B\ assms(2) calculation midpoint_bet) ultimately show ?thesis using OS_def by blast qed have "Col B B'' B'" proof - have "Coplanar A' B B'' B'" proof - have "Coplanar A' B' B B''" proof - have "\ Col A A' B'" using Col_perm \\ Col B' A' A\ by blast moreover have "Coplanar A A' B' B" using \A' B' OS A B\ ncoplanar_perm_8 os__coplanar by blast moreover have "Coplanar A A' B' B''" using \A' B' OS A B''\ ncoplanar_perm_8 os__coplanar by blast ultimately show ?thesis using coplanar_trans_1 by blast qed then show ?thesis using ncoplanar_perm_4 by blast qed moreover have "A' \ B'" by (simp add: P12) moreover have "Per B B' A'" by (simp add: P24 l8_2) moreover have "Per B'' B' A'" using Per_cases \Per A' B' B''\ by auto ultimately show ?thesis using cop_per2__col by blast qed have "Cong B B' A A'" using P6 P7 cong_inner_transitivity by blast have "B = B'' \ B' Midpoint B B''" proof - have "Col B B' B''" using \Col B B'' B'\ not_col_permutation_5 by blast moreover have "Cong B' B B' B''" by (metis Cong_perm P35 P6 P7 cong_inner_transitivity) ultimately show ?thesis using l7_20 by simp qed { assume "B = B''" then obtain M where S1: "Col X y M \ M Midpoint A B" using P30 by blast then have "R = M" using assms(4) l7_17 by auto have "A \ B" using assms(1) col_trivial_1 by auto have "Col R A B" by (simp add: assms(4) midpoint_col) have "X \ R" using Midpoint_def P28 \A' B' OS A B''\ \B = B''\ assms(4) midpoint_col one_side_chara by auto then have "\ X Y. (R PerpAt X Y A B \ X Y Perp P Q \ Coplanar A B C X \ Coplanar A B C Y)" proof - have "R PerpAt R X A B" proof - have "R X Perp A B" using P30 S1 \A \ B\ \B = B''\ \R = M\ \X \ R\ perp_col perp_left_comm by blast then show ?thesis using \Col R A B\ l8_14_2_1b_bis not_col_distincts by blast qed moreover have "R X Perp P Q" proof - have "X R Perp P Q" proof - have "X y Perp P Q" proof - have "P Q Perp X y" using P11 P13 P14 P29 P33A col_trivial_2 col_trivial_3 perp_col4 by blast then show ?thesis using Perp_perm by blast qed moreover have "Col X y R" by (simp add: S1 \R = M\) ultimately show ?thesis using \X \ R\ perp_col by blast qed then show ?thesis using Perp_perm by blast qed moreover have "Coplanar A B C R" using \Col R A B\ ncop__ncols not_col_permutation_2 by blast moreover have "Coplanar A B C X" proof - have "Col P Q X" using P12 P13 P14 P28A col3 by blast moreover have "\ Col P Q C" by (simp add: P2A) moreover have "Coplanar P Q C A" using assms(3) coplanar_perm_19 midpoint__coplanar by blast moreover have "Coplanar P Q C B" using assms(2) midpoint_col ncop__ncols not_col_permutation_5 by blast moreover have "Coplanar P Q C C" using ncop_distincts by auto moreover have "Coplanar P Q C X" using calculation(1) ncop__ncols by blast ultimately show ?thesis using coplanar_pseudo_trans by blast qed ultimately show ?thesis by blast qed } { assume "B' Midpoint B B''" have "A' B' TS B B''" proof - have "\ Col B A' B'" using Col_perm \\ Col A' B' B\ by blast moreover have "\ Col B'' A' B'" using \A' B' OS A B''\ col124__nos not_col_permutation_2 by blast moreover have "\ T. Col T A' B' \ Bet B T B''" using \B' Midpoint B B''\ col_trivial_3 midpoint_bet by blast ultimately show ?thesis by (simp add: TS_def) qed have "A' B' OS B B''" using \A' B' OS A B''\ \A' B' OS A B\ one_side_symmetry one_side_transitivity by blast have "\ A' B' OS B B''" using \A' B' TS B B''\ l9_9_bis by blast then have "False" by (simp add: \A' B' OS B B''\) then have "\ X Y. (R PerpAt X Y A B \ X Y Perp P Q \ Coplanar A B C X \ Coplanar A B C Y)" by auto } then show ?thesis using \B = B'' \ \X Y. R PerpAt X Y A B \ X Y Perp P Q \ Coplanar A B C X \ Coplanar A B C Y\ \B = B'' \ B' Midpoint B B''\ by blast qed lemma l13_1: assumes "\ Col A B C" and "P Midpoint B C" and "Q Midpoint A C" and "R Midpoint A B" shows "\ X Y.(R PerpAt X Y A B \ X Y Perp P Q)" proof - obtain X Y where "R PerpAt X Y A B \ X Y Perp P Q \ Coplanar A B C X \ Coplanar A B C Y" using l13_1_aux assms(1) assms(2) assms(3) assms(4) by blast then show ?thesis by blast qed lemma per_lt: assumes "A \ B" and "C \ B" and "Per A B C" shows "A B Lt A C \ C B Lt A C" proof - have "B A Lt A C \ B C Lt A C" using assms(1) assms(2) assms(3) l11_46 by auto then show ?thesis using lt_left_comm by blast qed lemma cong_perp_conga: assumes "Cong A B C B" and "A C Perp B P" shows "A B P CongA C B P \ B P TS A C" proof - have P1: "A \ C" using assms(2) perp_distinct by auto have P2: "B \ P" using assms(2) perp_distinct by auto have P3: "A \ B" by (metis P1 assms(1) cong_diff_3) have P4: "C \ B" using P3 assms(1) cong_diff by blast show ?thesis proof cases assume P5: "Col A B C" have P6: "\ Col B A P" using P3 P5 assms(2) col_transitivity_1 not_col_permutation_4 not_col_permutation_5 perp_not_col2 by blast have "Per P B A" using P3 P5 Perp_perm assms(2) not_col_permutation_5 perp_col1 perp_per_1 by blast then have P8: "Per A B P" using Per_cases by blast have "Per P B C" using P3 P5 P8 col_per2__per l8_2 l8_5 by blast then have P10: "Per C B P" using Per_perm by blast show ?thesis proof - have "A B P CongA C B P" using P2 P3 P4 P8 P10 l11_16 by auto moreover have "B P TS A C" by (metis Col_cases P1 P5 P6 assms(1) bet__ts between_cong not_cong_2143 not_cong_4321 third_point) ultimately show ?thesis by simp qed next assume T1: "\ Col A B C" obtain T where T2: "T PerpAt A C B P" using assms(2) perp_inter_perp_in by blast then have T3: "Col A C T \ Col B P T" using perp_in_col by auto have T4: "B \ T" using Col_perm T1 T3 by blast have T5: "B T Perp A C" using Perp_cases T3 T4 assms(2) perp_col1 by blast { assume T5_1: "A = T" have "B A Lt B C \ C A Lt B C" proof - have "B \ A" using P3 by auto moreover have "C \ A" using P1 by auto moreover have "Per B A C" using T5 T5_1 perp_comm perp_per_1 by blast ultimately show ?thesis by (simp add: per_lt) qed then have "False" using Cong_perm assms(1) cong__nlt by blast } then have T6: "A \ T" by auto { assume T6_1: "C = T" have "B C Lt B A \ A C Lt B A" proof - have "B \ C" using P4 by auto moreover have "A \ C" by (simp add: P1) moreover have "Per B C A" using T5 T6_1 perp_left_comm perp_per_1 by blast ultimately show ?thesis by (simp add: per_lt) qed then have "False" using Cong_perm assms(1) cong__nlt by blast } then have T7: "C \ T" by auto have T8: "T PerpAt B T T A" by (metis Perp_in_cases T2 T3 T4 T6 perp_in_col_perp_in) have T9: "T PerpAt B T T C" by (metis Col_cases T3 T7 T8 perp_in_col_perp_in) have T10: "Cong T A T C \ T A B CongA T C B \ T B A CongA T B C" proof - have "A T B CongA C T B" proof - have "Per A T B" using T2 perp_in_per_1 by auto moreover have "Per C T B" using T2 perp_in_per_3 by auto ultimately show ?thesis by (simp add: T4 T6 T7 l11_16) qed moreover have "Cong A B C B" by (simp add: assms(1)) moreover have "Cong T B T B" by (simp add: cong_reflexivity) moreover have "T B Le A B" proof - have "Per B T A" using T8 perp_in_per by auto then have "B T Lt B A \ A T Lt B A" using T4 T6 per_lt by blast then show ?thesis using Le_cases Lt_def by blast qed ultimately show ?thesis using l11_52 by blast qed show ?thesis proof - have T11: "A B P CongA C B P" proof - have "P B A CongA P B C" using Col_cases P2 T10 T3 col_conga__conga by blast thus ?thesis using conga_comm by blast qed moreover have "B P TS A C" proof - have T12: "A = C \ T Midpoint A C" using T10 T3 l7_20_bis not_col_permutation_5 by blast { assume "T Midpoint A C" then have "B P TS A C" by (smt Col_perm P2 T1 T3 \A = T \ False\ \C = T \ False\ col2__eq l9_18 midpoint_bet) } then show ?thesis using P1 T12 by auto qed ultimately show ?thesis by simp qed qed qed lemma perp_per_bet: assumes "\ Col A B C" and (* "Col A P C" and *) "Per A B C" and "P PerpAt P B A C" shows "Bet A P C" proof - have "A \ C" using assms(1) col_trivial_3 by auto then show ?thesis using assms(2) assms(3) l11_47 perp_in_left_comm by blast qed lemma ts_per_per_ts: assumes "A B TS C D" and "Per B C A" and "Per B D A" shows "C D TS A B" proof - have P1: "\ Col C A B" using TS_def assms(1) by blast have P2: "A \ B" using P1 col_trivial_2 by auto obtain P where P3: "Col P A B \ Bet C P D" using TS_def assms(1) by blast have P4: "C \ D" using assms(1) not_two_sides_id by auto show ?thesis proof - { assume "Col A C D" then have "C = D" by (metis assms(1) assms(2) assms(3) col_per2_cases col_permutation_2 not_col_distincts ts_distincts) then have "False" using P4 by auto } then have "\ Col A C D" by auto moreover have "\ Col B C D" using assms(1) assms(2) assms(3) per2_preserves_diff ts_distincts by blast moreover have "\ T. Col T C D \ Bet A T B" proof - have "Col P C D" using Col_def Col_perm P3 by blast moreover have "Bet A P B" proof - have "\ X. Col A B X \ A B Perp C X" using Col_perm P1 l8_18_existence by blast then obtain C' where P5: "Col A B C' \ A B Perp C C'" by blast have "\ X. Col A B X \ A B Perp D X" by (metis (no_types) Col_perm TS_def assms(1) l8_18_existence) then obtain D' where P6: "Col A B D' \ A B Perp D D'" by blast have P7: "A \ C'" using P5 assms(2) l8_7 perp_not_eq_2 perp_per_1 by blast have P8: "A \ D'" using P6 assms(3) l8_7 perp_not_eq_2 perp_per_1 by blast have P9: "Bet A C' B" proof - have "\ Col A C B" using Col_cases P1 by blast moreover have "Per A C B" by (simp add: assms(2) l8_2) moreover have "C' PerpAt C' C A B" using P5 Perp_in_perm l8_15_1 by blast ultimately show ?thesis using perp_per_bet by blast qed have P10: "Bet A D' B" proof - have "\ Col A D B" using P6 col_permutation_5 perp_not_col2 by blast moreover have "Per A D B" by (simp add: assms(3) l8_2) moreover have "D' PerpAt D' D A B" using P6 Perp_in_perm l8_15_1 by blast ultimately show ?thesis using perp_per_bet by blast qed show ?thesis proof cases assume "P = C'" then show ?thesis by (simp add: P9) next assume "P \ C'" show ?thesis proof cases assume "P = D'" then show ?thesis by (simp add: P10) next assume "P \ D'" show ?thesis proof cases assume "A = P" then show ?thesis by (simp add: between_trivial2) next assume "A \ P" show ?thesis proof cases assume "B = P" then show ?thesis using between_trivial by auto next assume "B \ P" have "Bet C' P D'" proof - have "Bet C P D" by (simp add: P3) moreover have "P \ C'" by (simp add: \P \ C'\) moreover have "P \ D'" by (simp add: \P \ D'\) moreover have "Col C' P D'" by (meson P2 P3 P5 P6 col3 col_permutation_2) moreover have "Per P C' C" using P3 P5 l8_16_1 l8_2 not_col_permutation_3 not_col_permutation_4 by blast moreover have "Per P D' D" by (metis P3 P6 calculation(3) not_col_permutation_2 perp_col2 perp_per_1) ultimately show ?thesis using per13_preserves_bet by blast qed then show ?thesis using P10 P9 bet3__bet by blast qed qed qed qed qed ultimately show ?thesis by auto qed ultimately show ?thesis by (simp add: TS_def) qed qed lemma l13_2_1: assumes "A B TS C D" and "Per B C A" and "Per B D A" and "Col C D E" and "A E Perp C D" and "C A B CongA D A B" shows "B A C CongA D A E \ B A D CongA C A E \ Bet C E D" proof - have P1: "\ Col C A B" using TS_def assms(1) by auto have P2: "A \ C" using P1 col_trivial_1 by blast have P3: "A \ B" using P1 col_trivial_2 by auto have P4: "A \ D" using assms(1) ts_distincts by auto have P5: "Cong B C B D \ Cong A C A D \ C B A CongA D B A" proof - have "\ Col B A C" by (simp add: P1 not_col_permutation_3) moreover have "A C B CongA A D B" using assms(1) assms(2) assms(3) l11_16 l8_2 ts_distincts by blast moreover have "B A C CongA B A D" by (simp add: assms(6) conga_comm) moreover have "Cong B A B A" by (simp add: cong_reflexivity) ultimately show ?thesis using l11_50_2 by blast qed then have P6: "C D Perp A B" using assms(1) assms(6) cong_conga_perp not_cong_2143 by blast then have P7: "C D TS A B" by (simp add: assms(1) assms(2) assms(3) ts_per_per_ts) obtain T1 where P8: "Col T1 C D \ Bet A T1 B" using P7 TS_def by auto obtain T where P9: "Col T A B \ Bet C T D" using TS_def assms(1) by blast have P10: "T1 = T" by (metis (no_types) Col_def P1 P3 P8 P9 between_equality_2 between_trivial2 l6_16_1) have P11: "T = E" proof - have "\ Col A B C" using Col_perm P1 by blast moreover have "C \ D" using assms(1) ts_distincts by blast moreover have "Col A B T" using Col_cases P9 by auto moreover have "Col A B E" by (metis P7 Perp_cases P6 assms(1) assms(5) col_perp2_ncol_col col_trivial_3 not_col_permutation_3 one_side_not_col123 os_ts1324__os ts_ts_os) moreover have "Col C D T" using NCol_cases P9 bet_col by blast moreover have "Col C D E" by (simp add: assms(4)) ultimately show ?thesis using l6_21 by blast qed show ?thesis proof - have "B A C CongA D A E" proof - have "A Out C C" using P2 out_trivial by auto moreover have "A Out B B" using P3 out_trivial by auto moreover have "A Out D D" using P4 out_trivial by auto moreover have "A Out E B" by (metis P10 P11 P7 P8 TS_def bet_out) ultimately show ?thesis by (meson assms(6) conga_comm conga_right_comm l11_10) qed moreover have "B A D CongA C A E" proof - have "C A E CongA D A B" by (meson Perp_cases P5 assms(5) assms(6) calculation cong_perp_conga conga_right_comm conga_trans not_cong_2143 not_conga_sym) then have "C A E CongA B A D" by (simp add: conga_right_comm) then show ?thesis by (simp add: conga_sym) qed moreover have "Bet C E D" using P11 P9 by auto ultimately show ?thesis by simp qed qed lemma triangle_mid_par: assumes "\ Col A B C" and "P Midpoint B C" and "Q Midpoint A C" shows "A B ParStrict Q P" proof - obtain R where P1: "R Midpoint A B" using midpoint_existence by auto then obtain X Y where P2: "R PerpAt X Y A B \ X Y Perp P Q \ Coplanar A B C X \ Coplanar A B C Y" using l13_1_aux assms(1) assms(2) assms(3) by blast have P3: "Coplanar X Y A P \ Coplanar X Y A Q \ Coplanar X Y B P \ Coplanar X Y B Q" proof - have "Coplanar A B C A" using ncop_distincts by auto moreover have "Coplanar A B C B" using ncop_distincts by auto moreover have "Coplanar A B C P" using assms(2) coplanar_perm_21 midpoint__coplanar by blast moreover have "Coplanar A B C Q" using assms(3) coplanar_perm_11 midpoint__coplanar by blast ultimately show ?thesis using P2 assms(1) coplanar_pseudo_trans by blast qed have P4: "Col X Y R \ Col A B R" using P2 perp_in_col by blast have P5: "R Y Perp A B \ X R Perp A B" using P2 perp_in_perp_bis by auto have P6: "Col A R B" using Col_perm P4 by blast have P7: "X \ Y" using P2 perp_not_eq_1 by auto { assume P8: "R Y Perp A B" have "Col Y R X" using P4 not_col_permutation_2 by blast then have "Y X Perp A B" using P2 Perp_cases perp_in_perp by blast then have P10: "X Y Perp A B" using Perp_cases by blast have "A B Par P Q" proof - have "Coplanar X Y A P" by (simp add: P3) moreover have "Coplanar X Y A Q" by (simp add: P3) moreover have "Coplanar X Y B P" by (simp add: P3) moreover have "Coplanar X Y B Q" by (simp add: P3) moreover have "A B Perp X Y" using P10 Perp_cases by auto moreover have "P Q Perp X Y" using P2 Perp_cases by auto ultimately show ?thesis using l12_9 by blast qed { assume "A B ParStrict P Q" then have "A B ParStrict Q P" using Par_strict_perm by blast } { assume "A \ B \ P \ Q \ Col A P Q \ Col B P Q" then have "Col A B P" using l6_16_1 not_col_permutation_1 by blast then have "P = B" by (metis Col_perm assms(1) assms(2) l6_16_1 midpoint_col) then have "A B ParStrict Q P" using assms(1) assms(2) col_trivial_2 is_midpoint_id by blast } then have "A B ParStrict Q P" using Par_def \A B Par P Q\ \A B ParStrict P Q \ A B ParStrict Q P\ by auto } { assume P10: "X R Perp A B" have "Col X R Y" by (simp add: Col_perm P4) then have P11: "X Y Perp A B" using P7 P10 perp_col by blast have "A B Par P Q" proof - have "A B Perp X Y" using P11 Perp_perm by blast moreover have "P Q Perp X Y" using P2 Perp_perm by blast ultimately show ?thesis using P3 l12_9 by blast qed { assume "A B ParStrict P Q" then have "A B ParStrict Q P" by (simp add: par_strict_right_comm) } { assume "A \ B \ P \ Q \ Col A P Q \ Col B P Q" then have "Col A B P" using Col_perm l6_16_1 by blast then have "P = B" by (metis Col_perm assms(1) assms(2) l6_16_1 midpoint_col) then have "A B ParStrict Q P" using assms(1) assms(2) col_trivial_2 is_midpoint_id by blast } then have "A B ParStrict Q P" using Par_def \A B Par P Q\ \A B ParStrict P Q \ A B ParStrict Q P\ by auto } then show ?thesis using P5 \R Y Perp A B \ A B ParStrict Q P\ by blast qed lemma cop4_perp_in2__col: assumes "Coplanar X Y A A'" and "Coplanar X Y A B'" and "Coplanar X Y B A'" and "Coplanar X Y B B'" and "P PerpAt A B X Y" and "P PerpAt A' B' X Y" shows "Col A B A'" proof - have P1: "Col A B P \ Col X Y P" using assms(5) perp_in_col by auto show ?thesis proof cases assume P2: "A = P" show ?thesis proof cases assume P3: "P = X" have "Col B A' P" proof - have "Coplanar Y B A' P" using P3 assms(3) ncoplanar_perm_18 by blast moreover have "Y \ P" using P3 assms(6) perp_in_distinct by blast moreover have "Per B P Y" using assms(5) perp_in_per_4 by auto moreover have "Per A' P Y" using assms(6) perp_in_per_2 by auto ultimately show ?thesis using cop_per2__col by auto qed then show ?thesis using Col_perm P2 by blast next assume P4: "P \ X" have "Col B A' P" proof - have "Coplanar X B A' P" by (metis P1 assms(3) assms(6) col2_cop__cop col_trivial_3 ncoplanar_perm_9 perp_in_distinct) moreover have "Per B P X" using assms(5) perp_in_per_3 by auto moreover have "Per A' P X" using assms(6) perp_in_per_1 by auto ultimately show ?thesis using cop_per2__col P4 by auto qed then show ?thesis using Col_perm P2 by blast qed next assume P5: "A \ P" have P6: "Per A P Y" using assms(5) perp_in_per_2 by auto show ?thesis proof cases assume P7: "P = A'" have P8: "Per B' P Y" using assms(6) perp_in_per_4 by auto have "Col A B' P" proof - have "Coplanar Y A B' P" using assms(2) by (metis P1 assms(6) col_transitivity_2 coplanar_trans_1 ncop__ncols perp_in_distinct) then show ?thesis using P6 P8 cop_per2__col by (metis assms(2) assms(5) assms(6) col_permutation_4 coplanar_perm_5 perp_in_distinct perp_in_per_1 perp_in_per_3) qed then show ?thesis using P1 P7 by auto next assume T1: "P \ A'" show ?thesis proof cases assume T2: "Y = P" { assume R1: "Coplanar X P A A' \ P PerpAt A B X P \ P PerpAt A' B' X P \ A \ P" then have R2: "Per A P X" using perp_in_per_1 by auto have "Per A' P X" using R1 perp_in_per_1 by auto then have "Col A B A'" by (metis R1 R2 PerpAt_def col_permutation_3 col_transitivity_2 cop_per2__col ncoplanar_perm_5) } then show ?thesis using P5 T1 T2 assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) by blast next assume P10: "Y \ P" have "Col A A' P" proof - have "Coplanar Y A A' P" by (metis P1 assms(1) assms(6) col2_cop__cop col_trivial_2 ncoplanar_perm_9 perp_in_distinct) moreover have "Per A P Y" by (simp add: P6) moreover have "Per A' P Y" using assms(6) perp_in_per_2 by auto ultimately show ?thesis using cop_per2__col P10 by auto qed then show ?thesis using P1 P5 col2__eq col_permutation_4 by blast qed qed qed qed lemma l13_2: assumes "A B TS C D" and "Per B C A" and "Per B D A" and "Col C D E" and "A E Perp C D" shows "B A C CongA D A E \ B A D CongA C A E \ Bet C E D" proof - have P2: "\ Col C A B" using TS_def assms(1) by auto have P3: "C \ D" using assms(1) not_two_sides_id by blast have P4: "\ C'. B A C CongA D A C' \ D A OS C' B" proof - have "\ Col B A C" using Col_cases P2 by auto moreover have "\ Col D A B" using TS_def assms(1) by blast ultimately show ?thesis by (simp add: angle_construction_1) qed then obtain E' where P5: "B A C CongA D A E' \ D A OS E' B" by blast have P6: "A \ B" using P2 not_col_distincts by blast have P7: "A \ C" using P2 not_col_distincts by blast have P8: "A \ D" using P5 os_distincts by blast have P9: "((A B TS C E' \ A E' TS D B) \ (A B OS C E' \ A E' OS D B \ C A B CongA D A E' \ B A E' CongA E' A B)) \ C A E' CongA D A B" by (metis P5 P6 conga_diff56 conga_left_comm conga_pseudo_refl l11_22) have P10: "C D TS A B" by (simp add: assms(1) assms(2) assms(3) ts_per_per_ts) have P11: "\ Col A C D" using P10 TS_def by auto obtain T where P12: "Col T A B \ Bet C T D" using TS_def assms(1) by blast obtain T2 where P13: "Col T2 C D \ Bet A T2 B" using P10 TS_def by auto then have P14: "T = T2" by (metis Col_def Col_perm P12 P2 P3 P6 l6_16_1) have P15: "B InAngle D A C" using P10 assms(1) l11_24 ts2__inangle by blast have P16: "C A B LeA C A D" by (simp add: P10 assms(1) inangle__lea ts2__inangle) have P17: "E' InAngle D A C" proof - have "D A E' LeA D A C" using P16 P5 P7 P8 conga_left_comm conga_pseudo_refl l11_30 by presburger moreover have "D A OS C E'" by (meson P11 P15 P5 col124__nos in_angle_one_side invert_one_side not_col_permutation_2 one_side_symmetry one_side_transitivity) ultimately show ?thesis by (simp add: lea_in_angle) qed obtain E'' where P18: "Bet D E'' C \ (E'' = A \ A Out E'' E')" using InAngle_def P17 by auto { assume "E'' = A" then have "B A C CongA D A E \ B A D CongA C A E \ Bet C E D" using Col_def P11 P18 by auto } { assume P19: "A Out E'' E'" then have P20: "B A C CongA D A E''" by (meson OS_def P5 Tarski_neutral_dimensionless.out2__conga Tarski_neutral_dimensionless_axioms col_one_side_out col_trivial_2 l9_18_R1 not_conga one_side_reflexivity) have P21: "A \ T" using P11 P13 P14 by auto have "B A C CongA D A E \ B A D CongA C A E \ Bet C E D" proof cases assume P22: "E'' = T" have P23: "C A B CongA D A B" proof - have "C A B CongA D A T" using P22 P20 conga_left_comm by blast moreover have "A Out C C" using P7 out_trivial by presburger moreover have "A Out B B" using P6 out_trivial by auto moreover have "A Out D D" using P8 out_trivial by auto moreover have "A Out B T" using Out_def P13 P14 P6 P21 by blast ultimately show ?thesis using l11_10 by blast qed then show ?thesis using assms(1) assms(2) assms(3) assms(4) assms(5) l13_2_1 by blast next assume P23A: "E'' \ T" have P24: "D \ E''" using P2 P20 col_trivial_3 ncol_conga_ncol not_col_permutation_3 by blast { assume P24A: "C = E''" have P24B: "C A OS B D" by (meson P10 assms(1) invert_one_side ts_ts_os) have P24C: "A Out B D" proof - have "C A B CongA C A D" using P20 P24A conga_comm by blast moreover have "C A OS B D" by (simp add: P24B) ultimately show ?thesis using conga_os__out by blast qed then have "False" using Col_def P5 one_side_not_col124 out_col by blast } then have P25: "C \ E''" by auto have P26: "A \ E''" using P19 out_diff1 by auto { assume "Col E'' A B" then have "E'' = T" by (smt P13 P14 P18 P2 P3 bet_col l6_21 not_col_permutation_2 not_col_permutation_3) then have "False" using P23A by auto } then have P27: "\ Col E'' A B" by auto have "(A B TS C E'' \ A E'' TS D B) \ (A B OS C E'' \ A E'' OS D B \ C A B CongA D A E'' \ B A E'' CongA E'' A B)" proof cases assume P27_0: "A B OS C E''" have "A E'' OS D B" proof - have P27_1: "A E'' TS D C" by (metis Col_def P10 P18 P24 TS_def P25 bet__ts invert_two_sides l6_16_1) moreover have "A E'' TS B C" proof - have "A E'' TS T C" proof - have "\ Col T A E''" by (metis NCol_cases P13 P14 P21 P27 bet_col col3 col_trivial_2) moreover have "\ Col C A E''" using P27_1 TS_def by auto moreover have "\ T0. (Col T0 A E'' \ Bet T T0 C)" by (meson P12 P18 P27_0 between_symmetry col_trivial_3 l5_3 one_side_chara) ultimately show ?thesis by (simp add: TS_def) qed moreover have "A Out T B" using Out_def P13 P14 P21 P6 by auto ultimately show ?thesis using col_trivial_1 l9_5 by blast qed ultimately show ?thesis using OS_def by auto qed thus ?thesis using P20 P27_0 conga_distinct conga_left_comm conga_pseudo_refl by blast next assume P27_2: "\ A B OS C E''" show ?thesis proof - have P27_3: "A B TS C E''" using P18 P2 P27_2 P27 assms(1) bet_cop__cop between_symmetry cop_nos__ts ts__coplanar by blast moreover have "A E'' TS D B" proof - have P27_3: "A B OS D E''" using P18 bet_ts__os between_symmetry calculation one_side_symmetry by blast have P27_4: "A E'' TS T D" proof - have "\ Col T A E''" by (metis NCol_cases P13 P14 P21 P27 bet_col col3 col_trivial_2) moreover have "\ Col D A E''" by (smt Col_def P11 P18 P24 P27_3 bet3__bet bet_col1 col3 col_permutation_5 col_two_sides_bet l5_1) moreover have "\ T0. (Col T0 A E'' \ Bet T T0 D)" by (meson Bet_perm P12 P18 P27_3 bet_col1 bet_out__bet between_exchange3 col_trivial_3 not_bet_out one_side_chara) ultimately show ?thesis by (simp add: TS_def) qed have "A E'' TS B D" proof - have "A E'' TS T D" using P27_4 by simp moreover have "Col A A E''" using col_trivial_1 by auto moreover have "A Out T B" using P13 P14 P21 bet_out by auto ultimately show ?thesis using l9_5 by blast qed thus ?thesis by (simp add: l9_2) qed ultimately show ?thesis by simp qed qed then have P28: "C A E'' CongA D A B" using l11_22 by (metis P20 P26 P6 conga_left_comm conga_pseudo_refl) obtain C' where P29: "Bet B C C' \ Cong C C' B C" using segment_construction by blast obtain D' where P30: "Bet B D D' \ Cong D D' B D" using segment_construction by blast have P31: "B A D Cong3 D' A D" proof - have "Per A D B" by (simp add: assms(3) l8_2) then obtain D'' where P31_2: "D Midpoint B D'' \ Cong A B A D''" using Per_def by auto have "D Midpoint B D'" using Cong_perm Midpoint_def P30 by blast then have "D' = D''" using P31_2 symmetric_point_uniqueness by auto thus ?thesis using Cong3_def Cong_perm P30 P31_2 cong_reflexivity by blast qed then have P32: "B A D CongA D' A D" using P6 P8 cong3_conga by auto have "B A C Cong3 C' A C" proof - obtain C'' where P33_1: "C Midpoint B C'' \ Cong A B A C''" using Per_def assms(2) l8_2 by blast have "C Midpoint B C'" using Cong_perm Midpoint_def P29 by blast then have "C' = C''" using P33_1 symmetric_point_uniqueness by auto thus ?thesis using Cong3_def Cong_perm P29 P33_1 cong_reflexivity by blast qed then have P34: "B A C CongA C' A C" using P6 P7 cong3_conga by auto have P35: "E'' A C' CongA D' A E''" proof - have "(A C TS E'' C' \ A D TS D' E'') \ (A C OS E'' C' \ A D OS D' E'')" proof - have P35_1: "C A OS D E''" by (metis Col_perm P11 P18 P25 bet_out between_symmetry one_side_symmetry out_one_side) have P35_2: "C A OS B D" using P10 assms(1) one_side_symmetry ts_ts_os by blast have P35_3: "C A TS B C'" by (metis P2 P29 bet__ts cong_diff_4 not_col_distincts) have P35_4: "C A OS B E''" using P35_1 P35_2 one_side_transitivity by blast have P35_5: "D A OS C E''" by (metis Col_perm P18 P24 P35_1 bet2__out l5_1 one_side_not_col123 out_one_side) have P35_6: "D A OS B C" by (simp add: P10 assms(1) invert_two_sides l9_2 one_side_symmetry ts_ts_os) have P35_7: "D A TS B D'" by (metis P30 TS_def assms(1) bet__ts cong_diff_3 ts_distincts) have P35_8: "D A OS B E''" using P35_5 P35_6 one_side_transitivity by blast have P35_9: "A C TS E'' C'" using P35_3 P35_4 invert_two_sides l9_8_2 by blast have "A D TS D' E''" using P35_7 P35_8 invert_two_sides l9_2 l9_8_2 by blast thus ?thesis using P35_9 by simp qed moreover have "E'' A C CongA D' A D" proof - have "E'' A C CongA B A D" by (simp add: P28 conga_comm) moreover have "B A D CongA D' A D" by (simp add: P32) ultimately show ?thesis using conga_trans by blast qed moreover have "C A C' CongA D A E''" proof - have "D A E'' CongA C A C'" proof - have "D A E'' CongA B A C" by (simp add: P20 conga_sym) moreover have "B A C CongA C A C'" by (simp add: P34 conga_right_comm) ultimately show ?thesis using conga_trans by blast qed thus ?thesis using not_conga_sym by blast qed ultimately show ?thesis using l11_22 by auto qed have P36: "D' \ B" using P30 assms(1) bet_neq32__neq ts_distincts by blast have P37: "C' \ B" using P29 assms(1) bet_neq32__neq ts_distincts by blast then have P38: "\ Col C' D' B" by (metis Col_def P10 P29 P30 P36 TS_def col_transitivity_2) have P39: "C' D' ParStrict C D" proof - have "\ Col C' D' B" by (simp add: P38) moreover have "D Midpoint D' B" using P30 l7_2 midpoint_def not_cong_3412 by blast moreover have "C Midpoint C' B" using P29 l7_2 midpoint_def not_cong_3412 by blast ultimately show ?thesis using triangle_mid_par by auto qed have P40: "A E'' TS C D" by (metis Bet_perm Col_def P10 P18 P24 TS_def \C = E'' \ False\ bet__ts col_transitivity_2 invert_two_sides) have P41: "B A TS C D" by (simp add: assms(1) invert_two_sides) have P42: "A B OS C C'" proof - have "\ Col A B C" by (simp add: P2 not_col_permutation_1) moreover have "Col A B B" by (simp add: col_trivial_2) moreover have "B Out C C'" by (metis P29 P37 bet_out cong_identity) ultimately show ?thesis using out_one_side_1 by blast qed have P43: "A B OS D D'" using out_one_side_1 by (metis Col_perm P30 TS_def assms(1) bet_out col_trivial_1) then have P44: "A B OS D D'" using invert_two_sides by blast have P45: "A B TS C' D" using P42 assms(1) l9_8_2 by blast then have P46: "A B TS C' D'" using P44 l9_2 l9_8_2 by blast have P47: "C' D' Perp A E''" proof - have "A E'' TS C' D'" proof - have "A Out C' D' \ E'' A TS C' D'" proof - have "E'' A C' CongA E'' A D'" by (simp add: P35 conga_right_comm) moreover have "Coplanar E'' A C' D'" proof - have f1: "B A OS C C'" by (metis P42 invert_one_side) have f2: "Coplanar B A C' C" by (meson P42 ncoplanar_perm_7 os__coplanar) have f3: "Coplanar D' A C' D" by (meson P44 P46 col124__nos coplanar_trans_1 invert_one_side ncoplanar_perm_7 os__coplanar ts__coplanar) have "Coplanar D' A C' C" using f2 f1 by (meson P46 col124__nos coplanar_trans_1 ncoplanar_perm_6 ncoplanar_perm_8 ts__coplanar) then show ?thesis using f3 by (meson P18 bet_cop2__cop ncoplanar_perm_6 ncoplanar_perm_7 ncoplanar_perm_8) qed ultimately show ?thesis using conga_cop__or_out_ts by simp qed then show ?thesis using P46 col_two_sides_bet invert_two_sides not_bet_and_out out_col by blast qed moreover have "Cong C' A D' A" using Cong3_def P31 \B A C Cong3 C' A C\ cong_inner_transitivity by blast moreover have "C' A E'' CongA D' A E''" by (simp add: P35 conga_left_comm) ultimately show ?thesis by (simp add: cong_conga_perp) qed have T1: "Cong A C' A D'" proof - have "Cong A C' A B" using Cong3_def Cong_perm \B A C Cong3 C' A C\ by blast moreover have "Cong A D' A B" using Cong3_def P31 not_cong_4321 by blast ultimately show ?thesis using Cong_perm \Cong A C' A B\ \Cong A D' A B\ cong_inner_transitivity by blast qed obtain R where T2: "R Midpoint C' D'" using midpoint_existence by auto have "\ X Y. (R PerpAt X Y C' D' \ X Y Perp D C \ Coplanar C' D' B X \ Coplanar C' D' B Y)" proof - have "\ Col C' D' B" by (simp add: P38) moreover have "D Midpoint D' B" using P30 l7_2 midpoint_def not_cong_3412 by blast moreover have "C Midpoint C' B" using Cong_perm Mid_perm Midpoint_def P29 by blast moreover have "R Midpoint C' D'" by (simp add: T2) ultimately show ?thesis using l13_1_aux by blast qed then obtain X Y where T3: "R PerpAt X Y C' D' \ X Y Perp D C \ Coplanar C' D' B X \ Coplanar C' D' B Y" by blast then have "X \ Y" using perp_not_eq_1 by blast have "C D Perp A E''" proof cases assume "A = R" then have W1: "A PerpAt C' D' A E''" using Col_def P47 T2 between_trivial2 l8_14_2_1b_bis midpoint_col by blast have "Coplanar B C' D' E''" proof - have "\ Col B C D" using P10 TS_def by auto moreover have "Coplanar B C D B" using ncop_distincts by auto moreover have "Coplanar B C D C'" using P29 bet_col ncop__ncols by blast moreover have "Coplanar B C D D'" using P30 bet_col ncop__ncols by blast moreover have "Coplanar B C D E''" by (simp add: P18 bet__coplanar coplanar_perm_22) ultimately show ?thesis using coplanar_pseudo_trans by blast qed have "Coplanar C' D' X E''" proof - have "\ Col B C' D'" by (simp add: P38 not_col_permutation_2) moreover have "Coplanar B C' D' X" using T3 ncoplanar_perm_8 by blast moreover have "Coplanar B C' D' E''" by (simp add: \Coplanar B C' D' E''\) ultimately show ?thesis using coplanar_trans_1 by blast qed have "Coplanar C' D' Y E''" proof - have "\ Col B C' D'" by (simp add: P38 not_col_permutation_2) moreover have "Coplanar B C' D' Y" by (simp add: T3 coplanar_perm_12) moreover have "Coplanar B C' D' E''" by (simp add: \Coplanar B C' D' E''\) ultimately show ?thesis using coplanar_trans_1 by blast qed have "Coplanar C' D' X A" proof - have "Col C' D' A" using T2 \A = R\ midpoint_col not_col_permutation_2 by blast moreover have "Col X A A" by (simp add: col_trivial_2) ultimately show ?thesis using ncop__ncols by blast qed have "Coplanar C' D' Y A" proof - have "Col C' D' A" using T2 \A = R\ midpoint_col not_col_permutation_2 by blast moreover have "Col Y A A" by (simp add: col_trivial_2) ultimately show ?thesis using ncop__ncols by blast qed have "Col X Y A" proof - have "Coplanar C' D' X A" by (simp add: \Coplanar C' D' X A\) moreover have "Coplanar C' D' X E''" by (simp add: \Coplanar C' D' X E''\) moreover have "Coplanar C' D' Y A" by (simp add: \Coplanar C' D' Y A\) moreover have "Coplanar C' D' Y E''" by (simp add: \Coplanar C' D' Y E''\) moreover have "A PerpAt X Y C' D'" using T3 \A = R\ Perp_in_cases by auto moreover have "A PerpAt A E'' C' D'" using Perp_in_cases \A PerpAt C' D' A E''\ by blast ultimately show ?thesis using cop4_perp_in2__col by blast qed have "Col X Y E''" proof - have "Coplanar C' D' X E''" using \Coplanar C' D' X E''\ by auto moreover have "Coplanar C' D' X A" by (simp add: \Coplanar C' D' X A\) moreover have "Coplanar C' D' Y E''" by (simp add: \Coplanar C' D' Y E''\) moreover have "Coplanar C' D' Y A" using \Coplanar C' D' Y A\ by auto moreover have "A PerpAt X Y C' D'" using T3 \A = R\ Perp_in_cases by auto moreover have "A PerpAt E'' A C' D'" using Perp_in_perm W1 by blast ultimately show ?thesis using cop4_perp_in2__col by blast qed have "A E'' Perp C D" proof cases assume "Y = A" show ?thesis proof - have "A \ E''" by (simp add: P26) moreover have "A X Perp C D" using T3 Perp_cases \Y = A\ by blast moreover have "Col A X E''" using Col_perm \Col X Y E''\ \Y = A\ by blast ultimately show ?thesis using perp_col by blast qed next assume "Y \ A" show ?thesis proof - have "A \ E''" by (simp add: P26) moreover have "A Y Perp C D" proof - have "Y X Perp C D" using T3 by (simp add: perp_comm) then have "Y A Perp C D" using \Col X Y A\ \Y \ A\ col_trivial_2 perp_col2 perp_left_comm by blast then show ?thesis using Perp_cases by blast qed moreover have "Col A Y E''" using Col_perm \Col X Y A\ \Col X Y E''\ \X \ Y\ col_transitivity_2 by blast ultimately show ?thesis using perp_col by blast qed qed thus ?thesis using Perp_perm by blast next assume "A \ R" have "R \ C'" using P46 T2 is_midpoint_id ts_distincts by blast have "Per A R C'" using T1 T2 Per_def by blast then have "R PerpAt A R R C'" by (simp add: \A \ R\ \R \ C'\ per_perp_in) then have "R PerpAt R C' A R" using Perp_in_perm by blast then have "R C' Perp A R \ R R Perp A R" using perp_in_perp by auto { assume "R C' Perp A R" then have "C' R Perp A R" by (simp add: \R C' Perp A R\ Perp_perm) have "C' D' Perp R A" by (metis P47 T2 \A \ R\ \Per A R C'\ \R \ C'\ col_per_perp midpoint_col perp_distinct perp_right_comm) then have "R PerpAt C' D' R A" using T2 l8_14_2_1b_bis midpoint_col not_col_distincts by blast have "Col B D D'" by (simp add: Col_def P30) have "Col B C C'" using Col_def P29 by auto have "Col D E'' C" using P18 bet_col by auto have "Col R C' D'" using \R PerpAt C' D' R A\ by (simp add: T2 midpoint_col) have "Col A E'' E'" by (simp add: P19 out_col) have "Coplanar C' D' X A" proof - have "\ Col B C' D'" using Col_perm P38 by blast moreover have "Coplanar B C' D' X" using T3 ncoplanar_perm_8 by blast moreover have "Coplanar B C' D' A" using P46 ncoplanar_perm_18 ts__coplanar by blast ultimately show ?thesis using coplanar_trans_1 by auto qed have "Coplanar C' D' Y A" proof - have "\ Col B C' D'" using Col_perm P38 by blast moreover have "Coplanar B C' D' Y" using T3 ncoplanar_perm_8 by blast moreover have "Coplanar B C' D' A" using P46 ncoplanar_perm_18 ts__coplanar by blast ultimately show ?thesis using coplanar_trans_1 by auto qed have "Coplanar C' D' X R" proof - have "Col C' D' R" using Col_perm \Col R C' D'\ by blast moreover have "Col X R R" by (simp add: col_trivial_2) ultimately show ?thesis using ncop__ncols by blast qed have "Coplanar C' D' Y R" using Col_perm T2 midpoint_col ncop__ncols by blast have "Col X Y A" proof - have "R PerpAt X Y C' D'" using T3 by simp moreover have "R PerpAt A R C' D'" using Perp_in_perm \R PerpAt C' D' R A\ by blast ultimately show ?thesis using \Coplanar C' D' Y R\ \Coplanar C' D' X R\ cop4_perp_in2__col \Coplanar C' D' X A\ \Coplanar C' D' Y A\ by blast qed have Z1: "Col X Y R" using T3 perp_in_col by blast have "Col A E'' R" proof - have "Coplanar C' D' E'' R" using Col_cases \Col R C' D'\ ncop__ncols by blast moreover have "A E'' Perp C' D'" using P47 Perp_perm by blast moreover have "A R Perp C' D'" using Perp_perm \C' D' Perp R A\ by blast ultimately show ?thesis using cop_perp2__col by blast qed then have "Col X Y E''" using Z1 by (metis (full_types) \A \ R\ \Col X Y A\ col_permutation_4 col_trivial_2 l6_21) have "Col A E'' R" proof - have "Coplanar C' D' E'' R" using Col_cases \Col R C' D'\ ncop__ncols by blast moreover have "A E'' Perp C' D'" using P47 Perp_perm by blast moreover have "A R Perp C' D'" using Perp_perm \C' D' Perp R A\ by blast ultimately show ?thesis using cop_perp2__col by blast qed have "Col A R X" using \Col X Y A\ \Col X Y R\ \X \ Y\ col_transitivity_1 not_col_permutation_3 by blast have "Col A R Y" using \Col X Y A\ \Col X Y R\ \X \ Y\ col_transitivity_2 not_col_permutation_3 by blast have "A E'' Perp C D" proof cases assume "X = A" show ?thesis proof - have "A \ E''" by (simp add: P26) moreover have "A Y Perp C D" using T3 \X = A\ perp_right_comm by blast moreover have "Col A Y E''" using Col_perm \A \ R\ \Col A E'' R\ \Col A R Y\ col_transitivity_1 by blast ultimately show ?thesis using perp_col by auto qed next assume "X \ A" show ?thesis proof - have "A X Perp C D" by (smt P3 T3 \Col X Y A\ \X \ A\ col_trivial_2 col_trivial_3 perp_col4) moreover have "Col A X E''" using Col_perm \A \ R\ \Col A E'' R\ \Col A R X\ col_transitivity_1 by blast ultimately show ?thesis using P26 perp_col by blast qed qed } { assume "R R Perp A R" then have "A E'' Perp C D" using perp_distinct by blast } then have "A E'' Perp C D" using Perp_cases \R C' Perp A R \ A E'' Perp C D\ \R C' Perp A R \ R R Perp A R\ by auto then show ?thesis using Perp_perm by blast qed show ?thesis proof - have "Col A E E''" proof - have "Coplanar C D E E'" using assms(4) col__coplanar by auto moreover have "A E Perp C D" using assms(5) by auto moreover have "A E'' Perp C D" using Perp_perm \C D Perp A E''\ by blast ultimately show ?thesis by (meson P11 col_perp2_ncol_col col_trivial_3 not_col_permutation_2) qed moreover have "E'' = E" proof - have f1: "C = E'' \ Col C E'' D" by (metis P18 bet_out_1 out_col) then have f2: "C = E'' \ Col C E'' E" using Col_perm P3 assms(4) col_transitivity_1 by blast have "\p. (C = E'' \ Col C p D) \ \ Col C E'' p" using f1 by (meson col_transitivity_1) then have "\p. \ Col E'' p A \ Col E'' E p" using f2 by (metis (no_types) Col_perm P11 assms(4)) then show ?thesis using Col_perm calculation col_transitivity_1 by blast qed ultimately show ?thesis by (metis Bet_perm P18 P20 P28 Tarski_neutral_dimensionless.conga_left_comm Tarski_neutral_dimensionless_axioms not_conga_sym) qed qed then have "B A C CongA D A E \ B A D CongA C A E \ Bet C E D" by blast } thus ?thesis using P18 \E'' = A \ B A C CongA D A E \ B A D CongA C A E \ Bet C E D\ by blast qed lemma perp2_refl: assumes "A \ B" shows "P Perp2 A B A B" proof cases assume "Col A B P" obtain X where "\ Col A B X" using assms not_col_exists by blast then obtain Q where "A B Perp Q P \ A B OS X Q" using \Col A B P\ l10_15 by blast thus ?thesis using Perp2_def Perp_cases col_trivial_3 by blast next assume "\ Col A B P" then obtain Q where "Col A B Q \ A B Perp P Q" using l8_18_existence by blast thus ?thesis using Perp2_def Perp_cases col_trivial_3 by blast qed lemma perp2_sym: assumes "P Perp2 A B C D" shows "P Perp2 C D A B" proof - obtain X Y where "Col P X Y \ X Y Perp A B \ X Y Perp C D" using Perp2_def assms by auto thus ?thesis using Perp2_def by blast qed lemma perp2_left_comm: assumes "P Perp2 A B C D" shows "P Perp2 B A C D" proof - obtain X Y where "Col P X Y \ X Y Perp A B \ X Y Perp C D" using Perp2_def assms by auto thus ?thesis using Perp2_def perp_right_comm by blast qed lemma perp2_right_comm: assumes "P Perp2 A B C D" shows "P Perp2 A B D C" proof - obtain X Y where "Col P X Y \ X Y Perp A B \ X Y Perp C D" using Perp2_def assms by auto thus ?thesis using Perp2_def perp_right_comm by blast qed lemma perp2_comm: assumes "P Perp2 A B C D" shows "P Perp2 B A D C" proof - obtain X Y where "Col P X Y \ X Y Perp A B \ X Y Perp C D" using Perp2_def assms by auto thus ?thesis using assms perp2_left_comm perp2_right_comm by blast qed lemma perp2_pseudo_trans: assumes "P Perp2 A B C D" and "P Perp2 C D E F" and "\ Col C D P" shows "P Perp2 A B E F" proof - obtain X Y where P1: "Col P X Y \ X Y Perp A B \ X Y Perp C D" using Perp2_def assms(1) by auto obtain X' Y' where P2: "Col P X' Y' \ X' Y' Perp C D \ X' Y' Perp E F" using Perp2_def assms(2) by auto have "X Y Par X' Y'" proof - have "Coplanar P C D X" proof cases assume "X = P" thus ?thesis using ncop_distincts by blast next assume "X \ P" then have "X P Perp C D" using Col_cases P1 perp_col by blast then have "Coplanar X P C D" by (simp add: perp__coplanar) thus ?thesis using ncoplanar_perm_18 by blast qed have "Coplanar P C D Y" proof cases assume "Y = P" thus ?thesis using ncop_distincts by blast next assume "Y \ P" then have "Y P Perp C D" by (metis (full_types) Col_cases P1 Perp_cases col_transitivity_2 perp_col2) then have "Coplanar Y P C D" by (simp add: perp__coplanar) thus ?thesis using ncoplanar_perm_18 by blast qed have "Coplanar P C D X'" proof cases assume "X' = P" thus ?thesis using ncop_distincts by blast next assume "X' \ P" then have "X' P Perp C D" using Col_cases P2 perp_col by blast then have "Coplanar X' P C D" by (simp add: perp__coplanar) thus ?thesis using ncoplanar_perm_18 by blast qed have "Coplanar P C D Y'" proof cases assume "Y' = P" thus ?thesis using ncop_distincts by blast next assume "Y' \ P" then have "Y' P Perp C D" by (metis (full_types) Col_cases P2 Perp_cases col_transitivity_2 perp_col2) then have "Coplanar Y' P C D" by (simp add: perp__coplanar) thus ?thesis using ncoplanar_perm_18 by blast qed show ?thesis proof - have "Coplanar C D X X'" using Col_cases \Coplanar P C D X'\ \Coplanar P C D X\ assms(3) coplanar_trans_1 by blast moreover have "Coplanar C D X Y'" using Col_cases \Coplanar P C D X\ \Coplanar P C D Y'\ assms(3) coplanar_trans_1 by blast moreover have "Coplanar C D Y X'" using Col_cases \Coplanar P C D X'\ \Coplanar P C D Y\ assms(3) coplanar_trans_1 by blast moreover have "Coplanar C D Y Y'" using Col_cases \Coplanar P C D Y'\ \Coplanar P C D Y\ assms(3) coplanar_trans_1 by blast ultimately show ?thesis using l12_9 P1 P2 by blast qed qed thus ?thesis proof - { assume "X Y ParStrict X' Y'" then have "Col X X' Y'" using P1 P2 \X Y ParStrict X' Y'\ par_not_col by blast } then have "Col X X' Y'" using Par_def \X Y Par X' Y'\ by blast moreover have "Col Y X' Y'" proof - { assume "X Y ParStrict X' Y'" then have "Col Y X' Y'" using P1 P2 \X Y ParStrict X' Y'\ par_not_col by blast } thus ?thesis using Par_def \X Y Par X' Y'\ by blast qed moreover have "X \ Y" using P1 perp_not_eq_1 by auto ultimately show ?thesis by (meson Perp2_def P1 P2 col_permutation_1 perp_col2) qed qed lemma col_cop_perp2__pars_bis: assumes "\ Col A B P" and "Col C D P" and "Coplanar A B C D" and "P Perp2 A B C D" shows "A B ParStrict C D" proof - obtain X Y where P1: "Col P X Y \ X Y Perp A B \ X Y Perp C D" using Perp2_def assms(4) by auto then have "Col X Y P" using Col_perm by blast obtain Q where "X \ Q \ Y \ Q \ P \ Q \ Col X Y Q" using \Col X Y P\ diff_col_ex3 by blast thus ?thesis by (smt P1 Perp_perm assms(1) assms(2) assms(3) col_cop_perp2_pars col_permutation_1 col_transitivity_2 not_col_distincts perp_col4 perp_distinct) qed lemma perp2_preserves_bet23: assumes "Bet PO A B" and "Col PO A' B'" and "\ Col PO A A'" and "PO Perp2 A A' B B'" shows "Bet PO A' B'" proof - have "A \ A'" using assms(3) not_col_distincts by auto show ?thesis proof cases assume "A' = B'" thus ?thesis using between_trivial by auto next assume "A' \ B'" { assume "A = B" then obtain X Y where P1: "Col PO X Y \ X Y Perp A A' \ X Y Perp A B'" using Perp2_def assms(4) by blast have "Col A A' B'" proof - have "Coplanar X Y A' B'" using Col_cases Coplanar_def P1 assms(2) by auto moreover have "A A' Perp X Y" using P1 Perp_perm by blast moreover have "A B' Perp X Y" using P1 Perp_perm by blast ultimately show ?thesis using cop_perp2__col by blast qed then have "False" using Col_perm \A' \ B'\ assms(2) assms(3) l6_16_1 by blast } then have "A \ B" by auto have "A A' Par B B'" proof - obtain X Y where P2: "Col PO X Y \ X Y Perp A A' \ X Y Perp B B'" using Perp2_def assms(4) by auto then have "Coplanar X Y A B" using Coplanar_def assms(1) bet_col not_col_permutation_2 by blast show ?thesis proof - have "Coplanar X Y A B'" by (metis (full_types) Col_cases P2 assms(2) assms(3) col_cop2__cop col_trivial_3 ncop__ncols perp__coplanar) moreover have "Coplanar X Y A' B" proof cases assume "Col A X Y" then have "Col Y X A" by (metis (no_types) Col_cases) then show ?thesis by (metis Col_cases P2 assms(1) assms(3) bet_col colx ncop__ncols not_col_distincts) next assume "\ Col A X Y" moreover have "Coplanar A X Y A'" using Coplanar_def P2 perp_inter_exists by blast moreover have "Coplanar A X Y B" using \Coplanar X Y A B\ ncoplanar_perm_8 by blast ultimately show ?thesis using coplanar_trans_1 by auto qed moreover have "Coplanar X Y A' B'" using Col_cases Coplanar_def P2 assms(2) by auto moreover have "A A' Perp X Y" using P2 Perp_perm by blast moreover have "B B' Perp X Y" using P2 Perp_perm by blast ultimately show ?thesis using \Coplanar X Y A B\ l12_9 by auto qed qed { assume "A A' ParStrict B B'" then have "A A' OS B B'" by (simp add: l12_6) have "A A' TS PO B" using Col_cases \A \ B\ assms(1) assms(3) bet__ts by blast then have "A A' TS B' PO" using \A A' OS B B'\ l9_2 l9_8_2 by blast then have "Bet PO A' B'" using Col_cases assms(2) between_symmetry col_two_sides_bet invert_two_sides by blast } thus ?thesis by (metis Col_cases Par_def \A A' Par B B'\ \A \ B\ assms(1) assms(3) bet_col col_trivial_3 l6_21) qed qed lemma perp2_preserves_bet13: assumes "Bet B PO C" and "Col PO B' C'" and "\ Col PO B B'" and "PO Perp2 B C' C B'" shows "Bet B' PO C'" proof cases assume "C' = PO" thus ?thesis using not_bet_distincts by blast next assume "C' \ PO" show ?thesis proof cases assume "B' = PO" thus ?thesis using between_trivial2 by auto next assume "B' \ PO" have "B \ PO" using assms(3) col_trivial_1 by auto have "Col B PO C" by (simp add: Col_def assms(1)) show ?thesis proof cases assume "B = C" thus ?thesis using \B = C\ \B \ PO\ assms(1) between_identity by blast next assume "B \ C" have "B C' Par C B'" proof - obtain X Y where P1: "Col PO X Y \ X Y Perp B C' \ X Y Perp C B'" using Perp2_def assms(4) by auto have "Coplanar X Y B C" by (meson P1 \Col B PO C\ assms(1) l9_18_R2 ncop__ncols not_col_permutation_2 not_col_permutation_5 ts__coplanar) have "Coplanar X Y C' B'" using Col_cases Coplanar_def P1 assms(2) by auto show ?thesis proof - have "Coplanar X Y B C" by (simp add: \Coplanar X Y B C\) moreover have "Coplanar X Y B B'" by (metis P1 \C' \ PO\ assms(1) assms(2) bet_cop__cop calculation col_cop2__cop not_col_permutation_5 perp__coplanar) moreover have "Coplanar X Y C' C" by (smt P1 \B \ PO\ \Col B PO C\ \Coplanar X Y C' B'\ assms(2) col2_cop__cop col_cop2__cop col_permutation_1 col_transitivity_2 coplanar_perm_1 perp__coplanar) moreover have "Coplanar X Y C' B'" by (simp add: \Coplanar X Y C' B'\) moreover have "B C' Perp X Y" using P1 Perp_perm by blast moreover have "C B' Perp X Y" by (simp add: P1 Perp_perm) ultimately show ?thesis using l12_9 by blast qed qed have "B C' ParStrict C B'" by (metis Out_def Par_def \B C' Par C B'\ \B \ C\ \B \ PO\ assms(1) assms(3) col_transitivity_1 not_col_permutation_4 out_col) have "B' \ PO" by (simp add: \B' \ PO\) obtain X Y where P5: "Col PO X Y \ X Y Perp B C' \ X Y Perp C B'" using Perp2_def assms(4) by auto have "X \ Y" using P5 perp_not_eq_1 by auto show ?thesis proof cases assume "Col X Y B" have "Col X Y C" using P5 \B \ PO\ \Col B PO C\ \Col X Y B\ col_permutation_1 colx by blast show ?thesis proof - have "Col B' PO C'" using Col_cases assms(2) by auto moreover have "Per PO C B'" by (metis P5 \B C' ParStrict C B'\ \Col X Y C\ assms(2) col_permutation_2 par_strict_not_col_2 perp_col2 perp_per_2) moreover have "Per PO B C'" using P5 \B \ PO\ \Col X Y B\ col_permutation_1 perp_col2 perp_per_2 by blast ultimately show ?thesis by (metis Tarski_neutral_dimensionless.per13_preserves_bet_inv Tarski_neutral_dimensionless_axioms \B C' ParStrict C B'\ assms(1) assms(3) between_symmetry not_col_distincts not_col_permutation_3 par_strict_not_col_2) qed next assume "\ Col X Y B" then obtain B0 where U1: "Col X Y B0 \ X Y Perp B B0" using l8_18_existence by blast have "\ Col X Y C" by (smt P5 \B C' ParStrict C B'\ \Col B PO C\ \\ Col X Y B\ assms(2) col_permutation_2 colx par_strict_not_col_2) then obtain C0 where U2: "Col X Y C0 \ X Y Perp C C0" using l8_18_existence by blast have "B0 \ PO" by (metis P5 Perp_perm \Col B PO C\ \Col X Y B0 \ X Y Perp B B0\ \\ Col X Y C\ assms(3) col_permutation_2 col_permutation_3 col_perp2_ncol_col) { assume "C0 = PO" then have "C PO Par C B'" by (metis P5 Par_def Perp_cases \Col X Y C0 \ X Y Perp C C0\ \\ Col X Y C\ col_perp2_ncol_col not_col_distincts not_col_permutation_3 perp_distinct) then have "False" by (metis \B C' ParStrict C B'\ assms(2) assms(3) col3 not_col_distincts par_id_2 par_strict_not_col_2) } then have "C0 \ PO" by auto have "Bet B0 PO C0" proof - have "Bet B PO C" by (simp add: assms(1)) moreover have "PO \ B0" using \B0 \ PO\ by auto moreover have "PO \ C0" using \C0 \ PO\ by auto moreover have "Col B0 PO C0" using U1 U2 P5 \X \ Y\ col3 not_col_permutation_2 by blast moreover have "Per PO B0 B" proof - have "B0 PerpAt PO B0 B0 B" proof cases assume "X = B0" have "B0 PO Perp B B0" by (metis P5 U1 calculation(2) col3 col_trivial_2 col_trivial_3 perp_col2) show ?thesis proof - have "B0 \ PO" using calculation(2) by auto moreover have "B0 Y Perp B B0" using U1 \X = B0\ by auto moreover have "Col B0 Y PO" using Col_perm P5 \X = B0\ by blast ultimately show ?thesis using \B0 PO Perp B B0\ perp_in_comm perp_perp_in by blast qed next assume "X \ B0" have "X B0 Perp B B0" using U1 \X \ B0\ perp_col by blast have "B0 PO Perp B B0" by (metis P5 U1 calculation(2) not_col_permutation_2 perp_col2) then have "B0 PerpAt B0 PO B B0" by (simp add: perp_perp_in) thus ?thesis using Perp_in_perm by blast qed then show ?thesis by (simp add: perp_in_per) qed moreover have "Per PO C0 C" proof - have "C0 PO Perp C C0" by (metis P5 U2 calculation(3) col3 col_trivial_2 col_trivial_3 perp_col2) then have "C0 PerpAt PO C0 C0 C" by (simp add: perp_in_comm perp_perp_in) thus ?thesis using perp_in_per_2 by auto qed ultimately show ?thesis using per13_preserves_bet by blast qed show ?thesis proof cases assume "C' = B0" have "B' = C0" proof - have "\ Col C' PO C" using P5 U1 \B0 \ PO\ \C' = B0\ \\ Col X Y C\ colx not_col_permutation_3 not_col_permutation_4 by blast moreover have "C \ C0" using U2 \\ Col X Y C\ by auto moreover have "Col C C0 B'" proof - have "Coplanar X Y C0 B'" proof - have "Col X Y C0" by (simp add: U2) moreover have "Col C0 B' C0" by (simp add: col_trivial_3) ultimately show ?thesis using ncop__ncols by blast qed moreover have "C C0 Perp X Y" using Perp_perm U2 by blast moreover have "C B' Perp X Y" using P5 Perp_perm by blast ultimately show ?thesis using cop_perp2__col by auto qed ultimately show ?thesis by (metis Col_def \C' = B0\ \Bet B0 PO C0\ assms(2) colx) qed show ?thesis using Bet_cases \B' = C0\ \C' = B0\ \Bet B0 PO C0\ by blast next assume "C' \ B0" then have "B' \ C0" by (metis P5 U1 U2 \C0 \ PO\ assms(2) col_permutation_1 colx l8_18_uniqueness) have "B C' Par B B0" proof - have "Coplanar X Y B B" using ncop_distincts by auto moreover have "Coplanar X Y B B0" using U1 ncop__ncols by blast moreover have "Coplanar X Y C' B" using P5 ncoplanar_perm_1 perp__coplanar by blast moreover have "Coplanar X Y C' B0" using \\ Col X Y B\ calculation(2) calculation(3) col_permutation_1 coplanar_perm_12 coplanar_perm_18 coplanar_trans_1 by blast moreover have "B C' Perp X Y" using P5 Perp_perm by blast moreover have "B B0 Perp X Y" using Perp_perm U1 by blast ultimately show ?thesis using l12_9 by blast qed { assume "B C' ParStrict B B0" have "Col B B0 C'" by (simp add: \B C' Par B B0\ par_id_3) } then have "Col B B0 C'" using \B C' Par B B0\ par_id_3 by blast have "Col C C0 B'" proof - have "Coplanar X Y C0 B'" by (simp add: U2 col__coplanar) moreover have "C C0 Perp X Y" by (simp add: Perp_perm U2) moreover have "C B' Perp X Y" using P5 Perp_perm by blast ultimately show ?thesis using cop_perp2__col by auto qed show ?thesis proof - have "Col B' PO C'" using assms(2) not_col_permutation_4 by blast moreover have "Per PO C0 B'" proof - have "C0 PerpAt PO C0 C0 B'" proof cases assume "X = C0" have "C0 PO Perp C B'" proof - have "C0 \ PO" by (simp add: \C0 \ PO\) moreover have "C0 Y Perp C B'" using P5 \X = C0\ by auto moreover have "Col C0 Y PO" using Col_perm P5 \X = C0\ by blast ultimately show ?thesis using perp_col by blast qed then have "B' C0 Perp C0 PO" using Perp_perm \B' \ C0\ \Col C C0 B'\ not_col_permutation_1 perp_col1 by blast then have "C0 PerpAt C0 B' PO C0" using Perp_perm perp_perp_in by blast thus ?thesis using Perp_in_perm by blast next assume "X \ C0" then have "X C0 Perp C B'" using P5 U2 perp_col by blast have "C0 PO Perp C B'" using Col_cases P5 U2 \C0 \ PO\ perp_col2 by blast then have "B' C0 Perp C0 PO" using Perp_cases \B' \ C0\ \Col C C0 B'\ col_permutation_2 perp_col by blast thus ?thesis using Perp_in_perm Perp_perm perp_perp_in by blast qed then show ?thesis using perp_in_per_2 by auto qed moreover have "Per PO B0 C'" proof - have "B0 PerpAt PO B0 B0 C'" proof - have "Col C' B B0" using Col_cases \Col B B0 C'\ by blast then have "C' B0 Perp X Y" using perp_col P5 Perp_cases \C' \ B0\ by blast show ?thesis proof - have "PO B0 Perp B0 C'" by (smt P5 U1 \B0 \ PO\ \C' \ B0\ \Col B B0 C'\ col_trivial_2 not_col_permutation_2 perp_col4) then show ?thesis using Perp_in_cases Perp_perm perp_perp_in by blast qed qed thus ?thesis by (simp add: perp_in_per) qed ultimately show ?thesis using \B0 \ PO\ \C0 \ PO\ \Bet B0 PO C0\ between_symmetry per13_preserves_bet_inv by blast qed qed qed qed qed qed lemma is_image_perp_in: assumes "A \ A'" and "X \ Y" and "A A' Reflect X Y" shows "\ P. P PerpAt A A' X Y" by (metis Perp_def Tarski_neutral_dimensionless.Perp_perm Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) ex_sym1 l10_6_uniqueness) lemma perp_inter_perp_in_n: assumes "A B Perp C D" shows "\ P. Col A B P \ Col C D P \ P PerpAt A B C D" by (simp add: assms perp_inter_perp_in) lemma perp2_perp_in: assumes "PO Perp2 A B C D" and "\ Col PO A B" and "\ Col PO C D" shows "\ P Q. Col A B P \ Col C D Q \ Col PO P Q \ P PerpAt PO P A B \ Q PerpAt PO Q C D" proof - obtain X Y where P1: "Col PO X Y \ X Y Perp A B \ X Y Perp C D" using Perp2_def assms(1) by blast have "X \ Y" using P1 perp_not_eq_1 by auto obtain P where P2: "Col X Y P \ Col A B P \ P PerpAt X Y A B" using P1 perp_inter_perp_in_n by blast obtain Q where P3: "Col X Y Q \ Col C D Q \ Q PerpAt X Y C D" using P1 perp_inter_perp_in_n by blast have "Col A B P" using P2 by simp moreover have "Col C D Q" using P3 by simp moreover have "Col PO P Q" using P2 P3 P1 \X \ Y\ col3 not_col_permutation_2 by blast moreover have "P PerpAt PO P A B" proof cases assume "X = PO" thus ?thesis by (metis P2 assms(2) not_col_permutation_3 not_col_permutation_4 perp_in_col_perp_in perp_in_sym) next assume "X \ PO" then have "P PerpAt A B X PO" by (meson Col_cases P1 P2 perp_in_col_perp_in perp_in_sym) then have "P PerpAt A B PO X" using Perp_in_perm by blast then have "P PerpAt A B PO P" by (metis Col_cases assms(2) perp_in_col perp_in_col_perp_in) thus ?thesis by (simp add: perp_in_sym) qed moreover have "Q PerpAt PO Q C D" by (metis P1 P3 \X \ Y\ assms(3) col_trivial_2 colx not_col_permutation_3 not_col_permutation_4 perp_in_col_perp_in perp_in_right_comm perp_in_sym) ultimately show ?thesis by blast qed lemma l13_8: assumes "U \ PO" and "V \ PO" and "Col PO P Q" and "Col PO U V" and "Per P U PO" and "Per Q V PO" shows "PO Out P Q \ PO Out U V" by (smt Out_def assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) l8_2 not_col_permutation_5 per23_preserves_bet per23_preserves_bet_inv per_distinct_1) lemma perp_in_rewrite: assumes "P PerpAt A B C D" shows "P PerpAt A P P C \ P PerpAt A P P D \ P PerpAt B P P C \ P PerpAt B P P D" by (metis assms per_perp_in perp_in_distinct perp_in_per_1 perp_in_per_3 perp_in_per_4) lemma perp_out_acute: assumes "B Out A C'" and "A B Perp C C'" shows "Acute A B C" proof - have "A \ B" using assms(1) out_diff1 by auto have "C' \ B" using Out_def assms(1) by auto then have "B C' Perp C C'" by (metis assms(1) assms(2) out_col perp_col perp_comm perp_right_comm) then have "Per C C' B" using Perp_cases perp_per_2 by blast then have "Acute C' C B \ Acute C' B C" by (metis \C' \ B\ assms(2) l11_43 perp_not_eq_2) have "C \ B" using \B C' Perp C C'\ l8_14_1 by auto show ?thesis proof - have "B Out A C'" by (simp add: assms(1)) moreover have "B Out C C" by (simp add: \C \ B\ out_trivial) moreover have "Acute C' B C" by (simp add: \Acute C' C B \ Acute C' B C\) ultimately show ?thesis using acute_out2__acute by auto qed qed lemma perp_bet_obtuse: assumes "B \ C'" and "A B Perp C C'" and "Bet A B C'" shows "Obtuse A B C" proof - have "Acute C' B C" proof - have "B Out C' C'" using assms(1) out_trivial by auto moreover have "Col A B C'" by (simp add: Col_def assms(3)) then have "C' B Perp C C'" using Out_def assms(2) assms(3) bet_col1 calculation perp_col2 by auto ultimately show ?thesis using perp_out_acute by blast qed thus ?thesis using acute_bet__obtuse assms(2) assms(3) between_symmetry perp_not_eq_1 by blast qed end subsubsection "Part 1: 2D" context Tarski_2D begin lemma perp_in2__col: assumes "P PerpAt A B X Y" and "P PerpAt A' B' X Y" shows "Col A B A'" using cop4_perp_in2__col all_coplanar assms by blast lemma perp2_trans: assumes "P Perp2 A B C D" and "P Perp2 C D E F" shows "P Perp2 A B E F" proof - obtain X Y where P1: "Col P X Y \ X Y Perp A B \ X Y Perp C D" using Perp2_def assms(1) by blast obtain X' Y' where P2: "Col P X' Y' \ X' Y' Perp C D \ X' Y' Perp E F" using Perp2_def assms(2) by blast { assume "X Y Par X' Y'" then have P3: "X Y ParStrict X' Y' \ (X \ Y \ X' \ Y' \ Col X X' Y' \ Col Y X' Y')" using Par_def by blast { assume "X Y ParStrict X' Y'" then have "P Perp2 A B E F" using P1 P2 par_not_col by auto } { assume "X \ Y \ X' \ Y' \ Col X X' Y' \ Col Y X' Y'" then have "P Perp2 A B E F" by (meson P1 P2 Perp2_def col_permutation_1 perp_col2) } then have "P Perp2 A B E F" using P3 \X Y ParStrict X' Y' \ P Perp2 A B E F\ by blast } { assume "\ X Y Par X' Y'" then have "P Perp2 A B E F" using P1 P2 l12_9_2D by blast } thus ?thesis using \X Y Par X' Y' \ P Perp2 A B E F\ by blast qed lemma perp2_par: assumes "PO Perp2 A B C D" shows "A B Par C D" using Perp2_def l12_9_2D Perp_perm assms by blast end subsubsection "Part 2: length" context Tarski_neutral_dimensionless begin lemma lg_exists: "\ l. (QCong l \ l A B)" using QCong_def cong_pseudo_reflexivity by blast lemma lg_cong: assumes "QCong l" and "l A B" and "l C D" shows "Cong A B C D" by (metis QCong_def assms(1) assms(2) assms(3) cong_inner_transitivity) lemma lg_cong_lg: assumes "QCong l" and "l A B" and "Cong A B C D" shows "l C D" by (metis QCong_def assms(1) assms(2) assms(3) cong_transitivity) lemma lg_sym: assumes "QCong l" and "l A B" shows "l B A" using assms(1) assms(2) cong_pseudo_reflexivity lg_cong_lg by blast lemma ex_points_lg: assumes "QCong l" shows "\ A B. l A B" using QCong_def assms cong_pseudo_reflexivity by fastforce lemma is_len_cong: assumes "TarskiLen A B l" and "TarskiLen C D l" shows "Cong A B C D" using TarskiLen_def assms(1) assms(2) lg_cong by auto lemma is_len_cong_is_len: assumes "TarskiLen A B l" and "Cong A B C D" shows "TarskiLen C D l" using TarskiLen_def assms(1) assms(2) lg_cong_lg by fastforce lemma not_cong_is_len: assumes "\ Cong A B C D" and "TarskiLen A B l" shows "\ l C D" using TarskiLen_def assms(1) assms(2) lg_cong by auto lemma not_cong_is_len1: assumes "\ Cong A B C D" and "TarskiLen A B l" shows "\ TarskiLen C D l" using assms(1) assms(2) is_len_cong by blast lemma lg_null_instance: assumes "QCongNull l" shows "l A A" by (metis QCongNull_def QCong_def assms cong_diff cong_trivial_identity) lemma lg_null_trivial: assumes "QCong l" and "l A A" shows "QCongNull l" using QCongNull_def assms(1) assms(2) by auto lemma lg_null_dec: (*assumes "QCong l" *) shows "QCongNull l \ \ QCongNull l" by simp lemma ex_point_lg: assumes "QCong l" shows "\ B. l A B" by (metis QCong_def assms not_cong_3412 segment_construction) lemma ex_point_lg_out: assumes "A \ P" and "QCong l" and "\ QCongNull l" shows "\ B. (l A B \ A Out B P)" proof - obtain X Y where P1: "\ X0 Y0. (Cong X Y X0 Y0 \ l X0 Y0)" using QCong_def assms(2) by auto then have "l X Y" using cong_reflexivity by auto then have "X \ Y" using assms(2) assms(3) lg_null_trivial by auto then obtain B where "A Out P B \ Cong A B X Y" using assms(1) segment_construction_3 by blast thus ?thesis using Cong_perm Out_cases P1 by blast qed lemma ex_point_lg_bet: assumes "QCong l" shows "\ B. (l M B \ Bet A M B)" proof - obtain X Y where P1: "\ X0 Y0. (Cong X Y X0 Y0 \ l X0 Y0)" using QCong_def assms by auto then have "l X Y" using cong_reflexivity by blast obtain B where "Bet A M B \ Cong M B X Y" using segment_construction by blast thus ?thesis using Cong_perm P1 by blast qed lemma ex_points_lg_not_col: assumes "QCong l" and "\ QCongNull l" shows "\ A B. (l A B \ \ Col A B P)" proof - have "\ B::'p. A \ B" using another_point by blast then obtain A::'p where "P \ A" by metis then obtain Q where "\ Col P A Q" using not_col_exists by auto then have "A \ Q" using col_trivial_2 by auto then obtain B where "l A B \ A Out B Q" using assms(1) assms(2) ex_point_lg_out by blast thus ?thesis by (metis \\ Col P A Q\ col_transitivity_1 not_col_permutation_1 out_col out_diff1) qed lemma ex_eql: assumes "\ A B. (TarskiLen A B l1 \ TarskiLen A B l2)" shows "l1 = l2" proof - obtain A B where P1: "TarskiLen A B l1 \ TarskiLen A B l2" using assms by auto have "\ A0 B0. (l1 A0 B0 \ l2 A0 B0)" by (metis TarskiLen_def \TarskiLen A B l1 \ TarskiLen A B l2\ lg_cong lg_cong_lg) have "\ A0 B0. (l1 A0 B0 \ l2 A0 B0)" proof - have "\ A0 B0. (l1 A0 B0 \ l2 A0 B0)" by (metis TarskiLen_def \TarskiLen A B l1 \ TarskiLen A B l2\ lg_cong lg_cong_lg) moreover have "\ A0 B0. (l2 A0 B0 \ l1 A0 B0)" by (metis TarskiLen_def \TarskiLen A B l1 \ TarskiLen A B l2\ lg_cong lg_cong_lg) ultimately show ?thesis by blast qed thus ?thesis by blast qed lemma all_eql: assumes "TarskiLen A B l1" and "TarskiLen A B l2" shows "l1 = l2" using assms(1) assms(2) ex_eql by auto lemma null_len: assumes "TarskiLen A A la" and "TarskiLen B B lb" shows "la = lb" by (metis TarskiLen_def all_eql assms(1) assms(2) lg_null_instance lg_null_trivial) lemma eqL_equivalence: assumes "QCong la" and "QCong lb" and "QCong lc" shows "la = la \ (la = lb \ lb = la) \ (la = lb \ lb = lc \ la = lc)" by simp lemma ex_lg: "\ l. (QCong l \ l A B)" by (simp add: lg_exists) lemma lg_eql_lg: assumes "QCong l1" and "l1 = l2" shows "QCong l2" using assms(1) assms(2) by auto lemma ex_eqL: assumes "QCong l1" and "QCong l2" and "\ A B. (l1 A B \ l2 A B)" shows "l1 = l2" using TarskiLen_def all_eql assms(1) assms(2) assms(3) by auto subsubsection "Part 3 : angles" lemma ang_exists: assumes "A \ B" and "C \ B" shows "\ a. (QCongA a \ a A B C)" proof - have "A B C CongA A B C" by (simp add: assms(1) assms(2) conga_refl) thus ?thesis using QCongA_def assms(1) assms(2) by auto qed lemma ex_points_eng: assumes "QCongA a" shows "\ A B C. (a A B C)" proof - obtain A B C where "A \ B \ C \ B \ (\ X Y Z. (A B C CongA X Y Z \ a X Y Z))" using QCongA_def assms by auto thus ?thesis using conga_pseudo_refl by blast qed lemma ang_conga: assumes "QCongA a" and "a A B C" and "a A' B' C'" shows "A B C CongA A' B' C'" proof - obtain A0 B0 C0 where "A0 \ B0 \ C0 \ B0 \ (\ X Y Z. (A0 B0 C0 CongA X Y Z \ a X Y Z))" using QCongA_def assms(1) by auto thus ?thesis by (meson assms(2) assms(3) not_conga not_conga_sym) qed lemma is_ang_conga: assumes "A B C Ang a" and "A' B' C' Ang a" shows "A B C CongA A' B' C'" using Ang_def ang_conga assms(1) assms(2) by auto lemma is_ang_conga_is_ang: assumes "A B C Ang a" and "A B C CongA A' B' C'" shows "A' B' C' Ang a" proof - have "QCongA a" using Ang_def assms(1) by auto then obtain A0 B0 C0 where "A0 \ B0 \ C0 \ B0 \ (\ X Y Z. (A0 B0 C0 CongA X Y Z \ a X Y Z))" using QCongA_def by auto thus ?thesis by (metis Ang_def assms(1) assms(2) not_conga) qed lemma not_conga_not_ang: assumes "QCongA a" and "\ A B C CongA A' B' C'" and "a A B C" shows "\ a A' B' C'" using ang_conga assms(1) assms(2) assms(3) by auto lemma not_conga_is_ang: assumes "\ A B C CongA A' B' C'" and "A B C Ang a" shows "\ a A' B' C'" using Ang_def ang_conga assms(1) assms(2) by auto lemma not_cong_is_ang1: assumes "\ A B C CongA A' B' C'" and "A B C Ang a" shows "\ A' B' C' Ang a" using assms(1) assms(2) is_ang_conga by blast lemma ex_eqa: assumes "\ A B C.(A B C Ang a1 \ A B C Ang a2)" shows "a1 = a2" proof - obtain A B C where P1: "A B C Ang a1 \ A B C Ang a2" using assms by auto { fix x y z assume "a1 x y z" then have "x y z Ang a1" using Ang_def assms by auto then have "x y z CongA A B C" using P1 not_cong_is_ang1 by blast then have "x y z Ang a2" using P1 is_ang_conga_is_ang not_conga_sym by blast then have "a2 x y z" using Ang_def assms by auto } { fix x y z assume "a2 x y z" then have "x y z Ang a2" using Ang_def assms by auto then have "x y z CongA A B C" using P1 not_cong_is_ang1 by blast then have "x y z Ang a1" using P1 is_ang_conga_is_ang not_conga_sym by blast then have "a1 x y z" using Ang_def assms by auto } then have "\ x y z. (a1 x y z) \ (a2 x y z)" using \\z y x. a1 x y z \ a2 x y z\ by blast then have "\x y. (\ z. (a1 x y) z = (a2 x y) z)" by simp then have "\ x y. (a1 x y) = (a2 x y)" using fun_eq_iff by auto thus ?thesis using fun_eq_iff by auto qed lemma all_eqa: assumes "A B C Ang a1" and "A B C Ang a2" shows "a1 = a2" using assms(1) assms(2) ex_eqa by blast lemma is_ang_distinct: assumes "A B C Ang a" shows "A \ B \ C \ B" using assms conga_diff1 conga_diff2 is_ang_conga by blast lemma null_ang: assumes "A B A Ang a1" and "C D C Ang a2" shows "a1 = a2" using all_eqa assms(1) assms(2) conga_trivial_1 is_ang_conga_is_ang is_ang_distinct by auto lemma flat_ang: assumes "Bet A B C" and "Bet A' B' C'" and "A B C Ang a1" and "A' B' C' Ang a2" shows "a1 = a2" proof - have "A B C Ang a2" proof - have "A' B' C' Ang a2" by (simp add: assms(4)) moreover have "A' B' C' CongA A B C" by (metis assms(1) assms(2) assms(3) calculation conga_line is_ang_distinct) ultimately show ?thesis using is_ang_conga_is_ang by blast qed then show ?thesis using assms(3) all_eqa by auto qed lemma ang_distinct: assumes "QCongA a" and "a A B C" shows "A \ B \ C \ B" proof - have "A B C Ang a" by (simp add: Ang_def assms(1) assms(2)) thus ?thesis using is_ang_distinct by auto qed lemma ex_ang: assumes "B \ A" and "B \ C" shows "\ a. (QCongA a \ a A B C)" using ang_exists assms(1) assms(2) by auto lemma anga_exists: assumes "A \ B" and "C \ B" and "Acute A B C" shows "\ a. (QCongAAcute a \ a A B C)" proof - have "A B C CongA A B C" by (simp add: assms(1) assms(2) conga_refl) thus ?thesis using assms(1) QCongAAcute_def assms(3) by blast qed lemma anga_is_ang: assumes "QCongAAcute a" shows "QCongA a" proof - obtain A0 B0 C0 where P1: "Acute A0 B0 C0 \ (\ X Y Z.(A0 B0 C0 CongA X Y Z \ a X Y Z))" using QCongAAcute_def assms by auto thus ?thesis using QCongA_def by (metis acute_distincts) qed lemma ex_points_anga: assumes "QCongAAcute a" shows "\ A B C. a A B C" by (simp add: anga_is_ang assms ex_points_eng) lemma anga_conga: assumes "QCongAAcute a" and "a A B C" and "a A' B' C'" shows "A B C CongA A' B' C'" by (meson Tarski_neutral_dimensionless.ang_conga Tarski_neutral_dimensionless_axioms anga_is_ang assms(1) assms(2) assms(3)) lemma is_anga_to_is_ang: assumes "A B C AngAcute a" shows "A B C Ang a" using AngAcute_def Ang_def anga_is_ang assms by auto lemma is_anga_conga: assumes "A B C AngAcute a" and "A' B' C' AngAcute a" shows "A B C CongA A' B' C'" using AngAcute_def anga_conga assms(1) assms(2) by auto lemma is_anga_conga_is_anga: assumes "A B C AngAcute a" and "A B C CongA A' B' C'" shows "A' B' C' AngAcute a" using Tarski_neutral_dimensionless.AngAcute_def Tarski_neutral_dimensionless.Ang_def Tarski_neutral_dimensionless.is_ang_conga_is_ang Tarski_neutral_dimensionless_axioms assms(1) assms(2) is_anga_to_is_ang by fastforce lemma not_conga_is_anga: assumes "\ A B C CongA A' B' C'" and "A B C AngAcute a" shows "\ a A' B' C'" using AngAcute_def anga_conga assms(1) assms(2) by auto lemma not_cong_is_anga1: assumes "\ A B C CongA A' B' C'" and "A B C AngAcute a" shows "\ A' B' C' AngAcute a" using assms(1) assms(2) is_anga_conga by auto lemma ex_eqaa: assumes "\ A B C. (A B C AngAcute a1 \ A B C AngAcute a2)" shows "a1 = a2" using all_eqa assms is_anga_to_is_ang by blast lemma all_eqaa: assumes "A B C AngAcute a1" and "A B C AngAcute a2" shows "a1 = a2" using assms(1) assms(2) ex_eqaa by blast lemma is_anga_distinct: assumes "A B C AngAcute a" shows "A \ B \ C \ B" using assms is_ang_distinct is_anga_to_is_ang by blast lemma null_anga: assumes "A B A AngAcute a1" and "C D C AngAcute a2" shows "a1 = a2" using assms(1) assms(2) is_anga_to_is_ang null_ang by blast lemma anga_distinct: assumes "QCongAAcute a" and "a A B C" shows "A \ B \ C \ B" using ang_distinct anga_is_ang assms(1) assms(2) by blast lemma out_is_len_eq: assumes "A Out B C" and "TarskiLen A B l" and "TarskiLen A C l" shows "B = C" using Out_def assms(1) assms(2) assms(3) between_cong not_cong_is_len1 by fastforce lemma out_len_eq: assumes "QCong l" and "A Out B C" and "l A B" and "l A C" shows "B = C" using out_is_len_eq using TarskiLen_def assms(1) assms(2) assms(3) assms(4) by auto lemma ex_anga: assumes "Acute A B C" shows "\ a. (QCongAAcute a \ a A B C)" using acute_distincts anga_exists assms by blast lemma not_null_ang_ang: assumes "QCongAnNull a" shows "QCongA a" using QCongAnNull_def assms by blast lemma not_null_ang_def_equiv: "QCongAnNull a \ (QCongA a \ (\ A B C. (a A B C \ \ B Out A C)))" proof - { assume "QCongAnNull a" have "QCongA a \ (\ A B C. (a A B C \ \ B Out A C))" using QCongAnNull_def \QCongAnNull a\ ex_points_eng by fastforce } { assume "QCongA a \ (\ A B C. (a A B C \ \ B Out A C))" have "QCongAnNull a" by (metis Ang_def QCongAnNull_def Tarski_neutral_dimensionless.l11_21_a Tarski_neutral_dimensionless_axioms \QCongA a \ (\A B C. a A B C \ \ B Out A C)\ not_conga_is_ang) } thus ?thesis using \QCongAnNull a \ QCongA a \ (\A B C. a A B C \ \ B Out A C)\ by blast qed lemma not_flat_ang_def_equiv: "QCongAnFlat a \ (QCongA a \ (\ A B C. (a A B C \ \ Bet A B C)))" proof - { assume "QCongAnFlat a" then have "QCongA a \ (\ A B C. (a A B C \ \ Bet A B C))" using QCongAnFlat_def ex_points_eng by fastforce } { assume "QCongA a \ (\ A B C. (a A B C \ \ Bet A B C))" have "QCongAnFlat a" proof - obtain pp :: 'p and ppa :: 'p and ppb :: 'p where f1: "QCongA a \ a pp ppa ppb \ \ Bet pp ppa ppb" using \QCongA a \ (\A B C. a A B C \ \ Bet A B C)\ by blast then have f2: "\p pa pb. pp ppa ppb CongA pb pa p \ \ a pb pa p" by (metis (no_types) Ang_def Tarski_neutral_dimensionless.not_cong_is_ang1 Tarski_neutral_dimensionless_axioms) then have f3: "\p pa pb. (Col pp ppa ppb \ \ a pb pa p) \ \ Bet pb pa p" by (metis (no_types) Col_def Tarski_neutral_dimensionless.ncol_conga_ncol Tarski_neutral_dimensionless_axioms) have f4: "\p pa pb. (\ Bet ppa ppb pp \ \ Bet pb pa p) \ \ a pb pa p" using f2 f1 by (metis Col_def Tarski_neutral_dimensionless.l11_21_a Tarski_neutral_dimensionless_axioms not_bet_and_out not_out_bet) have f5: "\p pa pb. (\ Bet ppb pp ppa \ \ Bet pb pa p) \ \ a pb pa p" using f2 f1 by (metis Col_def Tarski_neutral_dimensionless.l11_21_a Tarski_neutral_dimensionless_axioms not_bet_and_out not_out_bet) { assume "Bet ppa ppb pp" then have ?thesis using f4 f1 QCongAnFlat_def by blast } moreover { assume "Bet ppb pp ppa" then have ?thesis using f5 f1 QCongAnFlat_def by blast } ultimately show ?thesis using f3 f1 Col_def QCongAnFlat_def by blast qed } thus ?thesis using \QCongAnFlat a \ QCongA a \ (\A B C. a A B C \ \ Bet A B C)\ by blast qed lemma ang_const: assumes "QCongA a" and "A \ B" shows "\ C. a A B C" proof - obtain A0 B0 C0 where "A0 \ B0 \ C0 \ B0 \ (\ X Y Z. (A0 B0 C0 CongA X Y Z \ a X Y Z))" by (metis QCongA_def assms(1)) then have "(A0 B0 C0 CongA A0 B0 C0) \ a A0 B0 C0" by (simp add: conga_refl) then have "a A0 B0 C0" using \A0 \ B0 \ C0 \ B0 \ (\X Y Z. A0 B0 C0 CongA X Y Z \ a X Y Z)\ conga_refl by blast then show ?thesis using \A0 \ B0 \ C0 \ B0 \ (\X Y Z. A0 B0 C0 CongA X Y Z \ a X Y Z)\ angle_construction_3 assms(2) by blast qed lemma ang_sym: assumes "QCongA a" and "a A B C" shows "a C B A" proof - obtain A0 B0 C0 where "A0 \ B0 \ C0 \ B0 \ (\ X Y Z. (A0 B0 C0 CongA X Y Z \ a X Y Z))" by (metis QCongA_def assms(1)) then show ?thesis by (metis Tarski_neutral_dimensionless.ang_conga Tarski_neutral_dimensionless_axioms assms(1) assms(2) conga_left_comm conga_refl not_conga_sym) qed lemma ang_not_null_lg: assumes "QCongA a" and "QCong l" and "a A B C" and "l A B" shows "\ QCongNull l" by (metis QCongNull_def TarskiLen_def ang_distinct assms(1) assms(3) assms(4) cong_reverse_identity not_cong_is_len) lemma ang_distincts: assumes "QCongA a" and "a A B C" shows "A \ B \ C \ B" using ang_distinct assms(1) assms(2) by auto lemma anga_sym: assumes "QCongAAcute a" and "a A B C" shows "a C B A" by (simp add: ang_sym anga_is_ang assms(1) assms(2)) lemma anga_not_null_lg: assumes "QCongAAcute a" and "QCong l" and "a A B C" and "l A B" shows "\ QCongNull l" using ang_not_null_lg anga_is_ang assms(1) assms(2) assms(3) assms(4) by blast lemma anga_distincts: assumes "QCongAAcute a" and "a A B C" shows "A \ B \ C \ B" using anga_distinct assms(1) assms(2) by blast lemma ang_const_o: assumes "\ Col A B P" and "QCongA a" and "QCongAnNull a" and "QCongAnFlat a" shows "\ C. a A B C \ A B OS C P" proof - obtain A0 B0 C0 where P1: "A0 \ B0 \ C0 \ B0 \ (\ X Y Z. (A0 B0 C0 CongA X Y Z \ a X Y Z))" by (metis QCongA_def assms(2)) then have "a A0 B0 C0" by (simp add: conga_refl) then have T1: "A0 \ C0" using P1 Tarski_neutral_dimensionless.QCongAnNull_def Tarski_neutral_dimensionless_axioms assms(3) out_trivial by fastforce have "A \ B" using assms(1) col_trivial_1 by blast have "A0 \ B0 \ B0 \ C0" using P1 by auto then obtain C where P2: "A0 B0 C0 CongA A B C \ (A B OS C P \ Col A B C)" using angle_construction_2 assms(1) by blast then have "a A B C" by (simp add: P1) have P3: "A B OS C P \ Col A B C" using P2 by simp have P4: "\ A B C. (a A B C \ \ B Out A C)" using assms(3) by (simp add: QCongAnNull_def) have P5: "\ A B C. (a A B C \ \ Bet A B C)" using assms(4) QCongAnFlat_def by auto { assume "Col A B C" have "\ B Out A C" using P4 by (simp add: \a A B C\) have "\ Bet A B C" using P5 by (simp add: \a A B C\) then have "A B OS C P" using \Col A B C\ \\ B Out A C\ l6_4_2 by blast then have "\ C1. (a A B C1 \ A B OS C1 P)" using \a A B C\ by blast } then have "\ C1. (a A B C1 \ A B OS C1 P)" using P3 \a A B C\ by blast then show ?thesis by simp qed lemma anga_const: assumes "QCongAAcute a" and "A \ B" shows "\ C. a A B C" using Tarski_neutral_dimensionless.ang_const Tarski_neutral_dimensionless_axioms anga_is_ang assms(1) assms(2) by fastforce lemma null_anga_null_angaP: "QCongANullAcute a \ IsNullAngaP a" proof - have "QCongANullAcute a \ IsNullAngaP a" using IsNullAngaP_def QCongANullAcute_def ex_points_anga by fastforce moreover have "IsNullAngaP a \ QCongANullAcute a" by (metis IsNullAngaP_def QCongAnNull_def Tarski_neutral_dimensionless.QCongANullAcute_def Tarski_neutral_dimensionless_axioms anga_is_ang not_null_ang_def_equiv) ultimately show ?thesis by blast qed lemma is_null_anga_out: assumes (*"QCongAAcute a" and *) "a A B C" and "QCongANullAcute a" shows "B Out A C" using QCongANullAcute_def assms(1) assms(2) by auto lemma acute_not_bet: assumes "Acute A B C" shows "\ Bet A B C" using acute_col__out assms bet_col not_bet_and_out by blast lemma anga_acute: assumes "QCongAAcute a" and "a A B C" shows "Acute A B C" by (smt Tarski_neutral_dimensionless.QCongAAcute_def Tarski_neutral_dimensionless_axioms acute_conga__acute assms(1) assms(2)) lemma not_null_not_col: assumes "QCongAAcute a" and "\ QCongANullAcute a" and "a A B C" shows "\ Col A B C" proof - have "Acute A B C" using anga_acute assms(1) assms(3) by blast then show ?thesis using Tarski_neutral_dimensionless.IsNullAngaP_def Tarski_neutral_dimensionless_axioms acute_col__out assms(1) assms(2) assms(3) null_anga_null_angaP by blast qed lemma ang_cong_ang: assumes "QCongA a" and "a A B C" and "A B C CongA A' B' C'" shows "a A' B' C'" by (metis QCongA_def assms(1) assms(2) assms(3) not_conga) lemma is_null_ang_out: assumes (*"QCongA a" and *) "a A B C" and "QCongANull a" shows "B Out A C" proof - have "a A B C \ B Out A C" using QCongANull_def assms(2) by auto then show ?thesis by (simp add: assms(1)) qed lemma out_null_ang: assumes "QCongA a" and "a A B C" and "B Out A C" shows "QCongANull a" by (metis QCongANull_def QCongAnNull_def assms(1) assms(2) assms(3) not_null_ang_def_equiv) lemma bet_flat_ang: assumes "QCongA a" and "a A B C" and "Bet A B C" shows "AngFlat a" by (metis AngFlat_def QCongAnFlat_def assms(1) assms(2) assms(3) not_flat_ang_def_equiv) lemma out_null_anga: assumes "QCongAAcute a" and "a A B C" and "B Out A C" shows "QCongANullAcute a" using IsNullAngaP_def assms(1) assms(2) assms(3) null_anga_null_angaP by auto lemma anga_not_flat: assumes "QCongAAcute a" shows "QCongAnFlat a" by (metis (no_types, lifting) Tarski_neutral_dimensionless.QCongAnFlat_def Tarski_neutral_dimensionless.anga_is_ang Tarski_neutral_dimensionless_axioms assms bet_col is_null_anga_out not_bet_and_out not_null_not_col) lemma anga_const_o: assumes "\ Col A B P" and "\ QCongANullAcute a" and "QCongAAcute a" shows "\ C. (a A B C \ A B OS C P)" proof - have "QCongA a" by (simp add: anga_is_ang assms(3)) moreover have "QCongAnNull a" using QCongANullAcute_def assms(2) assms(3) calculation not_null_ang_def_equiv by auto moreover have "QCongAnFlat a" by (simp add: anga_not_flat assms(3)) ultimately show ?thesis by (simp add: ang_const_o assms(1)) qed lemma anga_conga_anga: assumes "QCongAAcute a" and "a A B C" and "A B C CongA A' B' C'" shows "a A' B' C'" using ang_cong_ang anga_is_ang assms(1) assms(2) assms(3) by blast lemma anga_out_anga: assumes "QCongAAcute a" and "a A B C" and "B Out A A'" and "B Out C C'" shows "a A' B C'" proof - have "A B C CongA A' B C'" by (simp add: assms(3) assms(4) l6_6 out2__conga) thus ?thesis using anga_conga_anga assms(1) assms(2) by blast qed lemma out_out_anga: assumes "QCongAAcute a" and "B Out A C" and "B' Out A' C'" and "a A B C" shows "a A' B' C'" proof - have "A B C CongA A' B' C'" by (simp add: assms(2) assms(3) l11_21_b) thus ?thesis using anga_conga_anga assms(1) assms(4) by blast qed lemma is_null_all: assumes "A \ B" and "QCongANullAcute a" shows "a A B A" proof - obtain A0 B0 C0 where "Acute A0 B0 C0 \ (\ X Y Z. (A0 B0 C0 CongA X Y Z \ a X Y Z))" using QCongAAcute_def QCongANullAcute_def assms(2) by auto then have "a A0 B0 C0" using acute_distincts conga_refl by blast thus ?thesis by (smt QCongANullAcute_def assms(1) assms(2) out_out_anga out_trivial) qed lemma anga_col_out: assumes "QCongAAcute a" and "a A B C" and "Col A B C" shows "B Out A C" proof - have "Acute A B C" using anga_acute assms(1) assms(2) by auto then have P1: "Bet A B C \ B Out A C" using acute_not_bet by auto then have "Bet C A B \ B Out A C" using assms(3) l6_4_2 by auto thus ?thesis using P1 assms(3) l6_4_2 by blast qed lemma ang_not_lg_null: assumes "QCong la" and "QCong lc" and "QCongA a" and "la A B" and "lc C B" and "a A B C" shows "\ QCongNull la \ \ QCongNull lc" by (metis ang_not_null_lg ang_sym assms(1) assms(2) assms(3) assms(4) assms(5) assms(6)) lemma anga_not_lg_null: assumes (*"QCong la" and "QCong lc" and*) "QCongAAcute a" and "la A B" and "lc C B" and "a A B C" shows "\ QCongNull la \ \ QCongNull lc" by (metis QCongNull_def anga_not_null_lg anga_sym assms(1) assms(2) assms(3) assms(4)) lemma anga_col_null: assumes "QCongAAcute a" and "a A B C" and "Col A B C" shows "B Out A C \ QCongANullAcute a" using anga_col_out assms(1) assms(2) assms(3) out_null_anga by blast lemma eqA_preserves_ang: assumes "QCongA a" and "a = b" shows "QCongA b" using assms(1) assms(2) by auto lemma eqA_preserves_anga: assumes "QCongAAcute a" and (* "QCongA b" and*) "a = b" shows "QCongAAcute b" using assms(1) assms(2) by auto section "Some postulates of the parallels" lemma euclid_5__original_euclid: assumes "Euclid5" shows "EuclidSParallelPostulate" proof - { fix A B C D P Q R assume P1: "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" obtain M where P2: "M Midpoint B C" using midpoint_existence by auto obtain D' where P3: "C Midpoint D D'" using symmetric_point_construction by auto obtain E where P4: "M Midpoint D' E" using symmetric_point_construction by auto have P5: "A \ B" using P1 os_distincts by blast have P6: "B \ C" using P1 os_distincts by blast have P7: "C \ D" using P1 os_distincts by blast have P10: "M \ B" using P2 P6 is_midpoint_id by auto have P11: "M \ C" using P2 P6 is_midpoint_id_2 by auto have P13: "C \ D'" using P3 P7 is_midpoint_id_2 by blast have P16: "\ Col B C A" using one_side_not_col123 P1 by blast have "B C OS D A" using P1 one_side_symmetry by blast then have P17: "\ Col B C D" using one_side_not_col123 P1 by blast then have P18: "\ Col M C D" using P2 Col_perm P11 col_transitivity_2 midpoint_col by blast then have P19: "\ Col M C D'" by (metis P13 P3 Col_perm col_transitivity_2 midpoint_col) then have P20: "\ Col D' C B" by (metis Col_perm P13 P17 P3 col_transitivity_2 midpoint_col) then have P21: "\ Col M C E" by (metis P19 P4 bet_col col2__eq col_permutation_4 midpoint_bet midpoint_distinct_2) have P22: "M C D' CongA M B E \ M D' C CongA M E B" using P13 l11_49 by (metis Cong_cases P19 P2 P4 l11_51 l7_13_R1 l7_2 midpoint_cong not_col_distincts) have P23: "Cong C D' B E" using P11 P2 P4 l7_13_R1 l7_2 by blast have P27: "C B TS D D'" by (simp add: P13 P17 P3 bet__ts midpoint_bet not_col_permutation_4) have P28: "A InAngle C B E" proof - have "C B A LeA C B E" proof - have "A B C LeA B C D'" proof - have "Bet D C D'" by (simp add: P3 midpoint_bet) then show ?thesis using P1 P7 P13 sams_chara by (metis sams_left_comm sams_sym) qed moreover have "A B C CongA C B A" using P5 P6 conga_pseudo_refl by auto moreover have "B C D' CongA C B E" by (metis CongA_def Mid_cases P2 P22 P4 P6 symmetry_preserves_conga) ultimately show ?thesis using l11_30 by blast qed moreover have "C B OS E A" proof - have "C B TS E D'" using P2 P20 P4 l7_2 l9_2 mid_two_sides not_col_permutation_1 by blast moreover have "C B TS A D'" using P27 \B C OS D A\ invert_two_sides l9_8_2 by blast ultimately show ?thesis using OS_def by blast qed ultimately show ?thesis using lea_in_angle by simp qed obtain A' where P30: "Bet C A' E \ (A' = B \ B Out A' A)" using P28 InAngle_def by auto { assume "A' = B" then have "Col D' C B" by (metis Col_def P2 P21 P30 P6 col_transitivity_1 midpoint_col) then have "False" by (simp add: P20) then have "\ Y. B Out A Y \ C Out D Y" by auto } { assume P31: "B Out A' A" have "\ I. BetS D' C I \ BetS B A' I" proof - have P32: "BetS B M C" using BetS_def Midpoint_def P10 P11 P2 by auto moreover have "BetS E M D'" using BetS_def Bet_cases P19 P21 P4 midpoint_bet not_col_distincts by fastforce moreover have "BetS C A' E" proof - have P32A: "C \ A'" using P16 P31 out_col by auto { assume "A' = E" then have P33: "B Out A E" using P31 l6_6 by blast then have "A B C B C D SumA D' C D" proof - have "D' C B CongA A B C" proof - have "D' C M CongA E B M" by (simp add: P22 conga_comm) moreover have "C Out D' D'" using P13 out_trivial by auto moreover have "C Out B M" using BetSEq Out_cases P32 bet_out_1 by blast moreover have "B Out A E" using P33 by auto moreover have "B Out C M" using BetSEq Out_def P32 by blast ultimately show ?thesis using l11_10 by blast qed moreover have "D' C B B C D SumA D' C D" by (simp add: P27 l9_2 ts__suma_1) moreover have "B C D CongA B C D" using P6 P7 conga_refl by auto moreover have "D' C D CongA D' C D" using P13 P7 conga_refl by presburger ultimately show ?thesis using conga3_suma__suma by blast qed then have "D' C D CongA P Q R" using P1 suma2__conga by auto then have "Bet P Q R" using Bet_cases P3 bet_conga__bet midpoint_bet by blast then have "False" using P1 by simp } then have "A' \ E" by auto then show ?thesis by (simp add: BetS_def P30 P32A) qed moreover have "\ Col B C D'" by (simp add: P20 not_col_permutation_3) moreover have "Cong B M C M" using Midpoint_def P2 not_cong_1243 by blast moreover have "Cong E M D' M" using Cong_perm Midpoint_def P4 by blast ultimately show ?thesis using euclid_5_def assms by blast qed then obtain Y where P34: "Bet D' C Y \ BetS B A' Y" using BetSEq by blast then have "\ Y. B Out A Y \ C Out D Y" proof - have P35: "B Out A Y" by (metis BetSEq Out_def P31 P34 l6_7) moreover have "C Out D Y" proof - have "D \ C" using P7 by auto moreover have "Y \ C" using P16 P35 l6_6 out_col by blast moreover have "D' \ C" using P13 by auto moreover have "Bet D C D'" by (simp add: P3 midpoint_bet) moreover have "Bet Y C D'" by (simp add: Bet_perm P34) ultimately show ?thesis using l6_2 by blast qed ultimately show ?thesis by auto qed } then have "\ Y. B Out A Y \ C Out D Y" using P30 \A' = B \ \Y. B Out A Y \ C Out D Y\ by blast } then show ?thesis using euclid_s_parallel_postulate_def by blast qed lemma tarski_s_euclid_implies_euclid_5: assumes "TarskiSParallelPostulate" shows "Euclid5" proof - { fix P Q R S T U assume P1: "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" have P1A: "BetS P T Q" using P1 by simp have P1B: "BetS R T S" using P1 by simp have P1C: "BetS Q U R" using P1 by simp have P1D: "\ Col P Q S" using P1 by simp have P1E: "Cong P T Q T" using P1 by simp have P1F: "Cong R T S T" using P1 by simp obtain V where P2: "P Midpoint R V" using symmetric_point_construction by auto have P3: "Bet V P R" using Mid_cases P2 midpoint_bet by blast then obtain W where P4: "Bet P W Q \ Bet U W V" using inner_pasch using BetSEq P1C by blast { assume "P = W" have "P \ V" by (metis BetSEq Bet_perm Col_def Cong_perm Midpoint_def P1A P1B P1D P1E P1F P2 between_trivial is_midpoint_id_2 l7_9) have "Col P Q S" proof - have f1: "Col V P R" by (meson Col_def P3) have f2: "Col U R Q" by (simp add: BetSEq Col_def P1) have f3: "Bet P T Q" using BetSEq P1 by fastforce have f4: "R = P \ Col V P U" by (metis (no_types) Col_def P4 \P = W\ \P \ V\ l6_16_1) have f5: "Col Q P T" using f3 by (meson Col_def) have f6: "Col T Q P" using f3 by (meson Col_def) have f7: "Col P T Q" using f3 by (meson Col_def) have f8: "Col P Q P" using Col_def P4 \P = W\ by blast have "Col R T S" by (meson BetSEq Col_def P1) then have "T = P \ Q = P" using f8 f7 f6 f5 f4 f2 f1 by (metis (no_types) BetSEq P1 \P \ V\ colx l6_16_1) then show ?thesis by (metis BetSEq P1) qed then have "False" by (simp add: P1D) } then have P5: "P \ W" by auto have "Bet V W U" using Bet_cases P4 by auto then obtain X Y where P7: "Bet P V X \ Bet P U Y \ Bet X Q Y" using assms(1) P1 P4 P5 tarski_s_parallel_postulate_def by blast have "Q S Par P R" proof - have "Q \ S" using P1D col_trivial_2 by auto moreover have "T Midpoint Q P" using BetSEq P1A P1E l7_2 midpoint_def not_cong_1243 by blast moreover have "T Midpoint S R" using BetSEq P1B P1F l7_2 midpoint_def not_cong_1243 by blast ultimately show ?thesis using l12_17 by auto qed then have P9: "Q S ParStrict P R" using P1D Par_def par_strict_symmetry par_symmetry by blast have P10: "Q S TS P Y" proof - have P10A: "P \ R" using P9 par_strict_distinct by auto then have P11: "P \ X" by (metis P2 P7 bet_neq12__neq midpoint_not_midpoint) have P12: "\ Col X Q S" proof - have "Q S ParStrict P R" by (simp add: P9) then have "Col P R X" by (metis P2 P3 P7 bet_col between_symmetry midpoint_not_midpoint not_col_permutation_4 outer_transitivity_between) then have "P X ParStrict Q S" using P9 Par_strict_perm P11 par_strict_col_par_strict by blast then show ?thesis using par_strict_not_col_2 by auto qed { assume W1: "Col Y Q S" have W2: "Q = Y" by (metis P12 P7 W1 bet_col bet_col1 colx) then have "\ Col Q P R" using P9 W1 par_not_col by auto then have W3: "Q = U" by (smt BetS_def Col_def P1C P7 W2 col_transitivity_2) then have "False" using BetS_def P1C by auto } then have "\ Col Y Q S" by auto then have "Q S TS X Y" by (metis P7 P12 bet__ts not_col_distincts not_col_permutation_1) moreover have "Q S OS X P" proof - have "P \ V" using P10A P2 is_midpoint_id_2 by blast then have "Q S ParStrict P X" by (meson Bet_perm P3 P7 P9 P11 bet_col not_col_permutation_4 par_strict_col_par_strict) then have "Q S ParStrict X P" by (simp add: par_strict_right_comm) then show ?thesis by (simp add: l12_6) qed ultimately show ?thesis using l9_8_2 by auto qed then obtain I where W4: "Col I Q S \ Bet P I Y" using TS_def by blast have "\ I. (BetS S Q I \ BetS P U I)" proof - have "BetS P U I" proof - have "P \ Y" using P10 not_two_sides_id by auto have W4A: "Bet P U I" proof - have W5: "Col P U I" using P7 W4 bet_col1 by auto { assume W6: "Bet U I P" have W7: "Q S OS P U" proof - have "Q S OS R U" proof - have "\ Col Q S R" using P9 par_strict_not_col_4 by auto moreover have "Q Out R U" using BetSEq Out_def P1C by blast ultimately show ?thesis by (simp add: out_one_side) qed moreover have "Q S OS P R" by (simp add: P9 l12_6) ultimately show ?thesis using one_side_transitivity by blast qed have W8: "I Out P U \ \ Col Q S P" by (simp add: P1D not_col_permutation_1) have "False" proof - have "I Out U P" using W4 W6 W7 between_symmetry one_side_chara by blast then show ?thesis using W6 not_bet_and_out by blast qed } { assume V1: "Bet I P U" have "P R OS I U" proof - have "P R OS I Q" proof - { assume "Q = I" then have "Col P Q S" by (metis BetSEq Col_def P1C P7 P9 V1 W4 between_equality outer_transitivity_between par_not_col) then have "False" using P1D by blast } then have "Q \ I" by blast moreover have "P R ParStrict Q S" using P9 par_strict_symmetry by blast moreover have "Col Q S I" using Col_cases W4 by blast ultimately show ?thesis using one_side_symmetry par_strict_all_one_side by blast qed moreover have "P R OS Q U" proof - have "Q S ParStrict P R" using P9 by blast have "R Out Q U \ \ Col P R Q" by (metis BetSEq Bet_cases Out_def P1C calculation col124__nos) then show ?thesis by (metis P7 V1 W4 \Bet U I P \ False\ between_equality col_permutation_2 not_bet_distincts out_col outer_transitivity_between) qed ultimately show ?thesis using one_side_transitivity by blast qed then have V2: "P Out I U" using P7 W4 bet2__out os_distincts by blast then have "Col P I U" using V1 not_bet_and_out by blast then have "False" using V1 V2 not_bet_and_out by blast } then moreover have "\ (Bet U I P \ Bet I P U)" using \Bet U I P \ False\ by auto ultimately show ?thesis using Col_def W5 by blast qed { assume "P = U" then have "Col P R Q" using BetSEq Col_def P1C by blast then have "False" using P9 par_strict_not_col_3 by blast } then have V6: "P \ U" by auto { assume "U = I" have "Q = U" proof - have f1: "BetS Q I R" using P1C \U = I\ by blast then have f2: "Col Q I R" using BetSEq Col_def by blast have f3: "Col I R Q" using f1 by (simp add: BetSEq Col_def) { assume "R \ Q" moreover { assume "(R \ Q \ R \ I) \ \ Col I Q R" moreover { assume "\p. (R \ Q \ \ Col I p I) \ Col Q I p" then have "I = Q" using f1 by (metis (no_types) BetSEq Col_def col_transitivity_2) } ultimately have "(\p pa. ((pa \ I \ \ Col pa p R) \ Col Q I pa) \ Col I pa p) \ I = Q" using f3 f2 by (metis (no_types) col_transitivity_2) } ultimately have "(\p pa. ((pa \ I \ \ Col pa p R) \ Col Q I pa) \ Col I pa p) \ I = Q" using f1 by (metis (no_types) BetSEq P9 W4 col_transitivity_2 par_strict_not_col_4) } then show ?thesis using f2 by (metis P9 W4 \U = I\ col_transitivity_2 par_strict_not_col_4) qed then have "False" using BetSEq P1C by blast } then have "U \ I" by auto then show ?thesis by (simp add: W4A V6 BetS_def) qed moreover have "BetS S Q I" proof - have "Q R TS S I" proof - have "Q R TS P I" proof - have "\ Col P Q R" using P9 col_permutation_5 par_strict_not_col_3 by blast moreover have "\ Col I Q R" proof - { assume "Col I Q R" then have "Col Q S R" proof - have f1: "\p pa pb. Col p pa pb \ \ BetS pb p pa" by (meson BetSEq Col_def) then have f2: "Col U I P" using \BetS P U I\ by blast have f3: "Col I P U" by (simp add: BetSEq Col_def \BetS P U I\) have f4: "\p. (U = Q \ Col Q p R) \ \ Col Q U p" by (metis BetSEq Col_def P1C col_transitivity_1) { assume "P \ Q" moreover { assume "(P \ Q \ U \ Q) \ Col Q P Q" then have "(P \ Q \ U \ Q) \ \ Col Q P R" using Col_cases \\ Col P Q R\ by blast moreover { assume "\p. ((U \ Q \ P \ Q) \ \ Col Q p P) \ Col Q P p" then have "U \ Q \ \ Col Q P P" by (metis col_transitivity_1) then have "\ Col U Q P" using col_transitivity_2 by blast } ultimately have "\ Col U Q P \ I \ Q" using f4 f3 by blast } ultimately have "I \ Q" using f2 f1 by (metis BetSEq P1C col_transitivity_1 col_transitivity_2) } then have "I \ Q" using BetSEq \BetS P U I\ by blast then show ?thesis by (simp add: W4 \Col I Q R\ col_transitivity_2) qed then have "False" using P9 par_strict_not_col_4 by blast } then show ?thesis by blast qed moreover have "Col U Q R" using BetSEq Bet_cases Col_def P1C by blast moreover have "Bet P U I" by (simp add: BetSEq \BetS P U I\) ultimately show ?thesis using TS_def by blast qed moreover have "Q R OS P S" proof - have "Q R Par P S" proof - have "Q \ R" using BetSEq P1 by blast moreover have "T Midpoint Q P" using BetSEq Bet_cases P1A P1E cong_3421 midpoint_def by blast moreover have "T Midpoint R S" using BetSEq P1B P1F midpoint_def not_cong_1243 by blast ultimately show ?thesis using l12_17 by blast qed then have "Q R ParStrict P S" by (simp add: P1D Par_def not_col_permutation_4) then show ?thesis using l12_6 by blast qed ultimately show ?thesis using l9_8_2 by blast qed then show ?thesis by (metis BetS_def W4 col_two_sides_bet not_col_permutation_2 ts_distincts) qed ultimately show ?thesis by auto qed } then show ?thesis using euclid_5_def by blast qed lemma tarski_s_implies_euclid_s_parallel_postulate: assumes "TarskiSParallelPostulate" shows "EuclidSParallelPostulate" by (simp add: assms euclid_5__original_euclid tarski_s_euclid_implies_euclid_5) theorem tarski_s_euclid_implies_playfair_s_postulate: assumes "TarskiSParallelPostulate" shows "PlayfairSPostulate" proof - { fix A1 A2 B1 B2 P C1 C2 assume P1: "\ Col P A1 A2 \ A1 A2 Par B1 B2 \ Col P B1 B2 \ A1 A2 Par C1 C2 \ Col P C1 C2" have P1A: "\ Col P A1 A2" by (simp add: P1) have P2: "A1 A2 Par B1 B2" by (simp add: P1) have P3: "Col P B1 B2" by (simp add: P1) have P4: "A1 A2 Par C1 C2" by (simp add: P1) have P5: "Col P C1 C2" by (simp add: P1) have P6: "A1 A2 ParStrict B1 B2" proof - have "A1 A2 Par B1 B2" by (simp add: P1) moreover have "Col B1 B2 P" using P3 not_col_permutation_2 by blast moreover have "\ Col A1 A2 P" by (simp add: P1A not_col_permutation_1) ultimately show ?thesis using par_not_col_strict by auto qed have P7: "A1 A2 ParStrict C1 C2" proof - have "A1 A2 Par C1 C2" by (simp add: P1) moreover have "Col C1 C2 P" using Col_cases P1 by blast moreover have "\ Col A1 A2 P" by (simp add: P1A not_col_permutation_1) ultimately show ?thesis using par_not_col_strict by auto qed { assume "\ Col C1 B1 B2 \ \ Col C2 B1 B2" have "\ C'. Col C1 C2 C' \ B1 B2 TS A1 C'" proof - have T2: "Coplanar A1 A2 P A1" using ncop_distincts by auto have T3: "Coplanar A1 A2 B1 B2" by (simp add: P1 par__coplanar) have T4: "Coplanar A1 A2 C1 C2" by (simp add: P7 pars__coplanar) have T5: "Coplanar A1 A2 P B1" using P1 col_trivial_2 ncop_distincts par__coplanar par_col2_par_bis by blast then have T6: "Coplanar A1 A2 P B2" using P3 T3 col_cop__cop by blast have T7: "Coplanar A1 A2 P C1" using P1 T4 col_cop__cop coplanar_perm_1 not_col_permutation_2 par_distincts by blast then have T8: "Coplanar A1 A2 P C2" using P5 T4 col_cop__cop by blast { assume "\ Col C1 B1 B2" moreover have "C1 \ C2" using P1 par_neq2 by auto moreover have "Col B1 B2 P" using P1 not_col_permutation_2 by blast moreover have "Col C1 C2 P" using Col_cases P5 by auto moreover have "\ Col B1 B2 C1" using Col_cases calculation(1) by auto moreover have "\ Col B1 B2 A1" using P6 par_strict_not_col_3 by auto moreover have "Coplanar B1 B2 C1 A1" using Col_cases P1A T5 T2 T6 T7 coplanar_pseudo_trans by blast ultimately have "\ C'. Col C1 C2 C' \ B1 B2 TS A1 C'" using cop_not_par_other_side by blast } { assume "\ Col C2 B1 B2" moreover have "C2 \ C1" using P1 par_neq2 by blast moreover have "Col B1 B2 P" using Col_cases P3 by auto moreover have "Col C2 C1 P" using Col_cases P5 by auto moreover have "\ Col B1 B2 C2" by (simp add: calculation(1) not_col_permutation_1) moreover have "\ Col B1 B2 A1" using P6 par_strict_not_col_3 by auto moreover have "Coplanar B1 B2 C2 A1" using Col_cases P1A T2 T5 T6 T8 coplanar_pseudo_trans by blast ultimately have "\ C'. Col C1 C2 C' \ B1 B2 TS A1 C'" using cop_not_par_other_side by (meson not_col_permutation_4) } then show ?thesis using \\ Col C1 B1 B2 \ \C'. Col C1 C2 C' \ B1 B2 TS A1 C'\ \\ Col C1 B1 B2 \ \ Col C2 B1 B2\ by blast qed then obtain C' where W1: "Col C1 C2 C' \ B1 B2 TS A1 C'" by auto then have W2: "\ Col A1 B1 B2" using TS_def by blast obtain B where W3: "Col B B1 B2 \ Bet A1 B C'" using TS_def W1 by blast obtain C where W4: "P Midpoint C' C" using symmetric_point_construction by blast then have W4A: "Bet A1 B C' \ Bet C P C'" using Mid_cases W3 midpoint_bet by blast then obtain D where W5: "Bet B D C \ Bet P D A1" using inner_pasch by blast have W6: "C' \ P" using P3 TS_def W1 by blast then have "A1 A2 Par C' P" by (meson P1 W1 not_col_permutation_2 par_col2_par) have W9: "A1 A2 ParStrict C' P" using Col_cases P5 P7 W1 W6 par_strict_col2_par_strict by blast then have W10: "B \ P" by (metis W6 W4A bet_out_1 out_col par_strict_not_col_3) have W11: "P \ C" using W6 W4 is_midpoint_id_2 by blast { assume "P = D" then have "False" by (metis Col_def P3 W1 W3 W4A W5 W10 W11 col_trivial_2 colx l9_18_R1) } then have "P \ D" by auto then obtain X Y where W12: "Bet P B X \ Bet P C Y \ Bet X A1 Y" using W5 assms tarski_s_parallel_postulate_def by blast then have "P \ X" using W10 bet_neq12__neq by auto then have "A1 A2 ParStrict P X" by (metis Col_cases P3 P6 W10 W12 W3 bet_col colx par_strict_col2_par_strict) then have W15: "A1 A2 OS P X" by (simp add: l12_6) have "P \ Y" using W11 W12 between_identity by blast then have "A1 A2 ParStrict P Y" by (metis Col_def W11 W12 W4A W9 col_trivial_2 par_strict_col2_par_strict) then have W16: "A1 A2 OS P Y" using l12_6 by auto have "Col A1 X Y" by (simp add: W12 bet_col col_permutation_4) then have "A1 Out X Y" using col_one_side_out W15 W16 using one_side_symmetry one_side_transitivity by blast then have "False" using W12 not_bet_and_out by blast } then have "Col C1 B1 B2 \ Col C2 B1 B2" by auto } { fix A1 A2 B1 B2 P C1 C2 assume P1: "Col P A1 A2 \ A1 A2 Par B1 B2 \ Col P B1 B2 \ A1 A2 Par C1 C2 \ Col P C1 C2" have "Col C1 B1 B2" by (smt P1 l9_10 not_col_permutation_3 not_strict_par2 par_col2_par par_comm par_id_5 par_symmetry ts_distincts) moreover have "Col C2 B1 B2" by (smt P1 l9_10 not_col_permutation_3 not_strict_par2 par_col2_par par_id_5 par_left_comm par_symmetry ts_distincts) ultimately have "Col C1 B1 B2 \ Col C2 B1 B2" by auto } then show ?thesis using playfair_s_postulate_def by (metis \\P C2 C1 B2 B1 A2 A1. \ Col P A1 A2 \ 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\) qed end end diff --git a/thys/LocalLexing/Ladder.thy b/thys/LocalLexing/Ladder.thy --- a/thys/LocalLexing/Ladder.thy +++ b/thys/LocalLexing/Ladder.thy @@ -1,3579 +1,3568 @@ theory Ladder imports TheoremD9 begin context LocalLexing begin definition LeftDerivationFix :: "sentence \ nat \ derivation \ nat \ sentence \ bool" where "LeftDerivationFix \ i D j \ = (is_sentence \ \ is_sentence \ \ LeftDerivation \ D \ \ i < length \ \ j < length \ \ \ ! i = \ ! j \ (\ E F. D = E@(derivation_shift F 0 (Suc j)) \ LeftDerivation (take i \) E (take j \) \ LeftDerivation (drop (Suc i) \) F (drop (Suc j) \)))" definition LeftDerivationIntro :: "sentence \ nat \ rule \ nat \ derivation \ nat \ sentence \ bool" where "LeftDerivationIntro \ i r ix D j \ = (\ \. LeftDerives1 \ i r \ \ ix < length (snd r) \ (snd r) ! ix = \ ! j \ LeftDerivationFix \ (i + ix) D j \)" lemma LeftDerivationFix_empty[simp]: "is_sentence \ \ i < length \ \ LeftDerivationFix \ i [] i \" apply (auto simp add: LeftDerivationFix_def) apply (rule_tac x="[]" in exI) apply auto done lemma Derive_empty[simp]: "Derive a [] = a" by (auto simp add: Derive_def) lemma LeftDerivation_append1: "LeftDerivation a (D@[(i, r)]) c \ \ b. LeftDerivation a D b \ LeftDerives1 b i r c" by (simp add: LeftDerivation_append) lemma Derivation_append1: "Derivation a (D@[(i, r)]) c \ \ b. Derivation a D b \ Derives1 b i r c" by (simp add: Derivation_append) lemma Derivation_take_derive: assumes "Derivation a D b" shows "Derivation a (take n D) (Derive a (take n D))" by (metis Derivation_append Derive append_take_drop_id assms) lemma LeftDerivation_take_derive: assumes "LeftDerivation a D b" shows "LeftDerivation a (take n D) (Derive a (take n D))" by (metis Derive LeftDerivation_append LeftDerivation_implies_Derivation append_take_drop_id assms) lemma Derivation_Derive_take_Derives1: assumes "N \ 0" assumes "N \ length D" assumes "Derivation a D b" assumes \: "\ = Derive a (take (N - 1) D)" assumes "\ = Derive a (take N D)" shows "Derives1 \ (fst (D ! (N - 1))) (snd (D ! (N - 1))) \" proof - let ?D1 = "take (N - 1) D" let ?D2 = "take N D" from assms have app: "?D2 = ?D1 @ [D ! (N - 1)]" apply auto by (metis Suc_less_eq Suc_pred le_imp_less_Suc take_Suc_conv_app_nth) from assms have "Derivation a ?D2 \" using Derivation_take_derive by blast with app show ?thesis using Derivation.simps Derivation_append Derive \ by auto qed lemma LeftDerivation_Derive_take_LeftDerives1: assumes "N \ 0" assumes "N \ length D" assumes "LeftDerivation a D b" assumes \: "\ = Derive a (take (N - 1) D)" assumes "\ = Derive a (take N D)" shows "LeftDerives1 \ (fst (D ! (N - 1))) (snd (D ! (N - 1))) \" proof - let ?D1 = "take (N - 1) D" let ?D2 = "take N D" from assms have app: "?D2 = ?D1 @ [D ! (N - 1)]" apply auto by (metis Suc_less_eq Suc_pred le_imp_less_Suc take_Suc_conv_app_nth) from assms have "LeftDerivation a ?D2 \" using LeftDerivation_take_derive by blast with app show ?thesis by (metis Derive LeftDerivation_append1 LeftDerivation_implies_Derivation \ prod.collapse) qed lemma LeftDerives1_skip_prefix: "length a \ i \ LeftDerives1 (a@b) i r (a@c) \ LeftDerives1 b (i - length a) r c" apply (auto simp add: LeftDerives1_def) using leftmost_skip_prefix apply blast by (simp add: Derives1_skip_prefix) lemma LeftDerives1_skip_suffix: assumes i: "i < length a" assumes D: "LeftDerives1 (a@c) i r (b@c)" shows "LeftDerives1 a i r b" proof - note Derives1_def[where u="a@c" and v="b@c" and i=i and r=r] then have "\x y N \. a @ c = x @ [N] @ y \ b @ c = x @ \ @ y \ is_sentence x \ is_sentence y \ (N, \) \ \ \ r = (N, \) \ i = length x" using D LeftDerives1_implies_Derives1 by auto then obtain x y N \ where split: "a @ c = x @ [N] @ y \ b @ c = x @ \ @ y \ is_sentence x \ is_sentence y \ (N, \) \ \ \ r = (N, \) \ i = length x" by blast from split have "length (a@c) = length (x @ [N] @ y)" by auto then have "length a + length c = length x + length y + 1" by simp with split have "length a + length c = i + length y + 1" by simp with i have len_c_y: "length c \ length y" by arith let ?y = "take (length y - length c) y" from split have ac: "a @ c = (x @ [N]) @ y" by auto note cancel_suffix[where a=a and c = c and b = "x@[N]" and d = "y", OF ac len_c_y] then have a: "a = x @ [N] @ ?y" by auto from split have bc: "b @ c = (x @ \) @ y" by auto note cancel_suffix[where a=b and c = c and b = "x@\" and d = "y", OF bc len_c_y] then have b: "b = x @ \ @ ?y" by auto from split len_c_y a b show ?thesis apply (simp only: LeftDerives1_def Derives1_def) apply (rule_tac conjI) using D LeftDerives1_def i leftmost_cons_less apply blast apply (rule_tac x=x in exI) apply (rule_tac x="?y" in exI) apply (rule_tac x="N" in exI) apply (rule_tac x="\" in exI) apply auto by (rule is_sentence_take) qed lemma LeftDerives1_X_is_part_of_rule[consumes 2, case_names Suffix Prefix]: assumes aXb: "LeftDerives1 \ i r (a@[X]@b)" assumes split: "splits_at \ i \ N \" assumes prefix: "\ \. \ = a @ [X] @ \ \ length a < i \ is_word (a @ [X]) \ LeftDerives1 \ (i - length a - 1) r b \ False" assumes suffix: "\ \. \ = \ @ [X] @ b \ LeftDerives1 \ i r a \ False" shows "\ u v. a = \ @ u \ b = v @ \ \ (snd r) = u@[X]@v" proof - have aXb_old: "Derives1 \ i r (a@[X]@b)" using LeftDerives1_implies_Derives1 aXb by blast have prefix_or: "is_prefix \ a \ is_proper_prefix a \" by (metis Derives1_prefix split aXb_old is_prefix_eq_proper_prefix) have is_word_\: "is_word \" using LeftDerives1_splits_at_is_word aXb assms(2) by blast have "is_proper_prefix a \ \ False" proof - assume proper:"is_proper_prefix a \" then have "\ u. u \ [] \ \ = a@u" by (metis is_proper_prefix_def) then obtain u where u: "u \ [] \ \ = a@u" by blast note splits_at = splits_at_\[OF aXb_old split] splits_at_combine[OF split] from splits_at have \1: "\ = take i \" by blast from splits_at have \2: "\ = take i (a@[X]@b)" by blast from splits_at have len\: "length \ = i" by blast with proper have lena: "length a < i" using append_eq_conv_conj drop_eq_Nil leI u by auto with is_word_\ \2 have is_word_aX: "is_word (a@[X])" by (simp add: is_word_terminals not_less take_Cons' u) from u \2 have "a@u = take i (a@[X]@b)" by auto with lena have "u = take (i - length a) ([X]@b)" by (simp add: less_or_eq_imp_le) with lena have uX: "u = [X]@(take (i - length a - 1) b)" by (simp add: not_less take_Cons') let ?\ = "(take (i - length a - 1) b) @ [N] @ \" from splits_at have f1: "\ = \ @ [N] @ \" by blast with u uX have f2: "\ = a @ [X] @ ?\" by simp note skip = LeftDerives1_skip_prefix[where a = "a @ [X]" and b = "?\" and r = r and i = i and c = b] then have D: "LeftDerives1 ?\ (i - length a - 1) r b" using One_nat_def Suc_leI aXb append_assoc diff_diff_left f2 lena length_Cons length_append length_append_singleton list.size(3) by fastforce note prefix[OF f2 lena is_word_aX D] then show "False" . qed with prefix_or have is_prefix: "is_prefix \ a" by blast from aXb have aXb': "LeftDerives1 \ i r ((a@[X])@b)" by auto then have aXb'_old: "Derives1 \ i r ((a@[X])@b)" by (simp add: LeftDerives1_implies_Derives1) note Derives1_suffix[OF aXb'_old split] then have suffix_or: "is_suffix \ b \ is_proper_suffix b \" by (metis is_suffix_eq_proper_suffix) have "is_proper_suffix b \ \ False" proof - assume proper: "is_proper_suffix b \" then have "\ u. u \ [] \ \ = u@b" by (metis is_proper_suffix_def) then obtain u where u: "u \ [] \ \ = u@b" by blast note splits_at = splits_at_\[OF aXb_old split] splits_at_combine[OF split] from splits_at have \1: "\ = drop (Suc i) \" by blast from splits_at have \2: "\ = drop (i + length (snd r)) (a @ [X] @ b)" by blast from splits_at have len\: "length \ = length \ - i - 1" by blast with proper have lenb: "length b < length \" by (metis is_proper_suffix_length_cmp) from u \2 have "u@b = drop (i + length (snd r)) ((a @ [X]) @ b)" by auto hence "u = drop (i + length (snd r)) (a @ [X])" by (metis drop_cancel_suffix) hence uX: "u = drop (i + length (snd r)) a @ [X]" by (metis drop_keep_last u) let ?\ = "\ @ [N] @ (drop (i + length (snd r)) a)" from splits_at have f1: "\ = \ @ [N] @ \" by blast with u uX have f2: "\ = ?\ @ [X] @ b" by simp note skip = LeftDerives1_skip_suffix[where a = "?\" and c = "[X]@b" and b="a" and r = r and i = i] have f3: "i < length (\ @ [N] @ drop (i + length (snd r)) a)" proof - have f1: "1 + i + length b = length [X] + length b + i" by (metis Groups.add_ac(2) Suc_eq_plus1_left length_Cons list.size(3) list.size(4) semiring_normalization_rules(22)) have f2: "length \ - i - 1 = length ((\ @ [N] @ drop (i + length (snd r)) a) @ [X] @ b) - Suc i" by (metis f2 length_drop splits_at(1)) have "length ([]::symbol list) \ length \ - i - 1 - length b" by (metis add_diff_cancel_right' append_Nil2 append_eq_append_conv len\ length_append u) then have "length ([]::symbol list) \ length \ + length ([N] @ drop (i + length (snd r)) a) - i" using f2 f1 by (metis Suc_eq_plus1_left add_diff_cancel_right' diff_diff_left length_append) then show ?thesis by auto qed from aXb f2 have D: "LeftDerives1 (?\ @ [X] @ b) i r (a@[X]@b)" by auto note skip[OF f3 D] note suffix[OF f2 skip[OF f3 D]] then show "False" . qed with suffix_or have is_suffix: "is_suffix \ b" by blast from is_prefix have "\ u. a = \ @ u" by (auto simp add: is_prefix_def) then obtain u where u: "a = \ @ u" by blast from is_suffix have "\ v. b = v @ \" by (auto simp add: is_suffix_def) then obtain v where v: "b = v @ \" by blast from u v splits_at_combine[OF split] aXb have D:"LeftDerives1 (\@[N]@\) i r (\@(u@[X]@v)@\)" by simp from splits_at_\[OF aXb_old split] have i: "length \ = i" by blast from i have i1: "length \ \ i" and i2: "i \ length \" by auto note LeftDerives1_skip_suffix[OF _ LeftDerives1_skip_prefix[OF i1 D], simplified, OF i2] then have "LeftDerives1 [N] 0 r (u @ [X] @ v)" by auto then have "Derives1 [N] 0 r (u @ [X] @ v)" using LeftDerives1_implies_Derives1 by auto then have r: "snd r = u @ [X] @ v" by (metis Derives1_split append_Cons append_Nil length_0_conv list.inject self_append_conv) show ?thesis using u v r by auto qed lemma LeftDerivationFix_grow_suffix: assumes LDF: "LeftDerivationFix (b1@[X]@b2) (length b1) D j c" assumes suffix_b2: "LeftDerives1 suffix e r b2" assumes is_word_b1X: "is_word (b1@[X])" shows "LeftDerivationFix (b1@[X]@suffix) (length b1) ((e + length (b1@[X]), r)#D) j c" proof - from LDF have LDF': "is_sentence (b1@[X]@b2) \ is_sentence c \ LeftDerivation (b1 @ [X] @ b2) D c \ length b1 < length (b1 @ [X] @ b2) \ j < length c \ (b1 @ [X] @ b2) ! length b1 = c ! j \ (\E F. D = E @ derivation_shift F 0 (Suc j) \ LeftDerivation (take (length b1) (b1 @ [X] @ b2)) E (take j c) \ LeftDerivation (drop (Suc (length b1)) (b1 @ [X] @ b2)) F (drop (Suc j) c))" using LeftDerivationFix_def by blast then obtain E F where EF: "D = E @ derivation_shift F 0 (Suc j) \ LeftDerivation (take (length b1) (b1 @ [X] @ b2)) E (take j c) \ LeftDerivation (drop (Suc (length b1)) (b1 @ [X] @ b2)) F (drop (Suc j) c)" by blast then have LD_b1_c: "LeftDerivation b1 E (take j c)" by simp with is_word_b1X have E: "E = []" using LeftDerivation_implies_Derivation is_word_Derivation is_word_append by blast then have b1_def: "b1 = take j c" using LD_b1_c by auto then have b1_len: "j = length b1" by (simp add: LDF' dual_order.strict_implies_order min.absorb2) have D: "D = derivation_shift F 0 (Suc j)" using EF E by simp have step: "LeftDerives1 (b1 @ [X] @ suffix) (Suc (e + length b1)) r (b1 @ [X] @ b2) \ LeftDerivation (b1 @ [X] @ b2) D c" by (metis LDF' LeftDerives1_append_prefix add_Suc_right append_assoc assms(2) is_word_b1X length_append_singleton) then have is_sentence_b1Xsuffix: "is_sentence (b1 @ [X] @ suffix)" using Derives1_sentence1 LeftDerives1_implies_Derives1 by blast have X_eq_cj: "X = c ! j" using LDF' by auto show ?thesis apply (simp add: LeftDerivationFix_def) apply (rule conjI) using is_sentence_b1Xsuffix apply simp apply (rule conjI) using LDF' apply simp apply (rule conjI) using step apply force apply (rule conjI) using LDF' apply simp apply (rule conjI) apply (rule X_eq_cj) apply (rule_tac x="[]" in exI) apply (rule_tac x="(e, r)#F" in exI) apply auto apply (rule b1_len[symmetric]) apply (rule D) apply (rule b1_def) apply (rule_tac x="b2" in exI) apply (simp add: suffix_b2) using EF by auto qed lemma Derives1_append_suffix: assumes Derives1: "Derives1 v i r w" assumes u: "is_sentence u" shows "Derives1 (v@u) i r (w@u)" proof - have "\ \ N \. splits_at v i \ N \" using assms splits_at_ex by auto then obtain \ N \ where split_v: "splits_at v i \ N \" by blast have split_w: "w = \@(snd r)@\" using assms split_v splits_at_combine_dest by blast have split_uv: "splits_at (v@u) i \ N (\@u)" by (simp add: split_v splits_at_append) have is_sentence_uv: "is_sentence (v@u)" using Derives1 Derives1_sentence1 is_sentence_concat u by blast show ?thesis by (metis Derives1 Derives1_nonterminal Derives1_rule append_assoc is_sentence_uv split_uv split_v split_w splits_at_implies_Derives1) qed lemma leftmost_append_suffix: "leftmost i v \ leftmost i (v@u)" by (simp add: leftmost_def nth_append) lemma LeftDerives1_append_suffix: assumes Derives1: "LeftDerives1 v i r w" assumes u: "is_sentence u" shows "LeftDerives1 (v@u) i r (w@u)" proof - have 1: "Derives1 v i r w" by (simp add: Derives1 LeftDerives1_implies_Derives1) have 2: "leftmost i v" using Derives1 LeftDerives1_def by blast have 3: "is_sentence u" using u by fastforce have 4: "Derives1 (v@u) i r (w@u)" by (simp add: "1" "3" Derives1_append_suffix) have 5: "leftmost i (v@u)" by (simp add: "2" leftmost_append_suffix u) show ?thesis by (simp add: "4" "5" LeftDerives1_def) qed lemma LeftDerivationFix_is_sentence: "LeftDerivationFix a i D j b \ is_sentence a \ is_sentence b" using LeftDerivationFix_def by blast lemma LeftDerivationIntro_is_sentence: "LeftDerivationIntro \ i r ix D j \ \ is_sentence \ \ is_sentence \" by (meson Derives1_sentence1 LeftDerivationFix_is_sentence LeftDerivationIntro_def LeftDerives1_implies_Derives1) lemma LeftDerivationFix_grow_prefix: assumes LDF: "LeftDerivationFix (b1@[X]@b2) (length b1) D j c" assumes prefix_b1: "LeftDerives1 prefix e r b1" shows "LeftDerivationFix (prefix@[X]@b2) (length prefix) ((e, r)#D) j c" proof - from LDF have LDF': "LeftDerivation (b1 @ [X] @ b2) D c \ length b1 < length (b1 @ [X] @ b2) \ j < length c \ (b1 @ [X] @ b2) ! length b1 = c ! j \ (\E F. D = E @ derivation_shift F 0 (Suc j) \ LeftDerivation (take (length b1) (b1 @ [X] @ b2)) E (take j c) \ LeftDerivation (drop (Suc (length b1)) (b1 @ [X] @ b2)) F (drop (Suc j) c))" using LeftDerivationFix_def by blast then obtain E F where EF: "D = E @ derivation_shift F 0 (Suc j) \ LeftDerivation (take (length b1) (b1 @ [X] @ b2)) E (take j c) \ LeftDerivation (drop (Suc (length b1)) (b1 @ [X] @ b2)) F (drop (Suc j) c)" by blast then have E_b1_c: "LeftDerivation b1 E (take j c)" by simp with EF have F_b2_c: "LeftDerivation b2 F (drop (Suc j) c)" by simp have step: "LeftDerives1 (prefix @ [X] @ b2) e r (b1 @ [X] @ b2)" using LDF LeftDerivationFix_is_sentence LeftDerives1_append_suffix is_sentence_concat prefix_b1 by blast show ?thesis apply (simp add: LeftDerivationFix_def) apply (rule conjI) apply (metis Derives1_sentence1 LDF LeftDerivationFix_def LeftDerives1_implies_Derives1 is_sentence_concat is_sentence_cons prefix_b1) apply (rule conjI) using LDF LeftDerivationFix_is_sentence apply blast apply (rule conjI) apply (rule_tac x="b1@[X]@b2" in exI) using step apply simp using LDF' apply auto[1] apply (rule conjI) using LDF' apply simp apply (rule conjI) using LDF' apply auto[1] apply (rule_tac x="(e,r)#E" in exI) apply (rule_tac x="F" in exI) apply (auto simp add: EF F_b2_c) apply (rule_tac x="b1" in exI) apply (simp add: prefix_b1 E_b1_c) done qed lemma LeftDerivationFixOrIntro: "LeftDerivation a D \ \ is_sentence \ \ j < length \ \ (\ i. LeftDerivationFix a i D j \) \ (\ d \ ix. d < length D \ LeftDerivation a (take d D) \ \ LeftDerivationIntro \ (fst (D ! d)) (snd (D ! d)) ix (drop (Suc d) D) j \)" proof (induct "length D" arbitrary: a D \ j rule: less_induct) (* The induction here is unnecessary, but we use it anyway for context reasons *) case less have "length D = 0 \ length D \ 0" by blast then show ?case proof (induct rule: disjCases2) case 1 then have D: "D = []" by auto with less have "\i. LeftDerivationFix a i D j \" apply (rule_tac x=j in exI) by auto then show ?case by blast next case 2 note less2 = 2 have "\ n \ i. n \ length D \ \ = Derive a (take n D) \ LeftDerivationFix \ i (drop n D) j \" apply (rule_tac x="length D" in exI) apply auto using Derive LeftDerivationFix_empty LeftDerivation_implies_Derivation less by blast then show ?case proof (induct rule: ex_minimal_witness) case (Minimal N) then obtain \ i where Minimal_N: "N \ length D \ \ = Derive a (take N D) \ LeftDerivationFix \ i (drop N D) j \" by blast have "N = 0 \ N \ 0" by blast then show ?case proof (induct rule: disjCases2) case 1 with Minimal_N have "\ = a" by auto with 1 Minimal_N show ?case apply (rule_tac disjI1) by auto next case 2 let ?\ = "Derive a (take (N - 1) D)" have LeftDerives1_\: "LeftDerives1 ?\ (fst (D ! (N - 1))) (snd (D ! (N - 1))) \" using "2.hyps" LeftDerivation_Derive_take_LeftDerives1 Minimal_N less.prems(1) by blast then have Derives1_\: "Derives1 ?\ (fst (D ! (N - 1))) (snd (D ! (N - 1))) \" using LeftDerives1_implies_Derives1 by blast have i_len: "i < length \" using Minimal_N by (auto simp add: LeftDerivationFix_def) then have "\ X \_1 \_2. splits_at \ i \_1 X \_2" using splits_at_def by blast then obtain X \_1 \_2 where \_split: "splits_at \ i \_1 X \_2" by blast then have \_combine: "\ = \_1 @ [X] @ \_2" using splits_at_combine by blast then have LeftDerives1_\_hyp: "LeftDerives1 ?\ (fst (D ! (N - 1))) (snd (D ! (N - 1))) (\_1 @ [X] @ \_2)" using LeftDerives1_\ by blast from \_split have i_def: "i = length \_1" by (simp add: dual_order.strict_implies_order min.absorb2 splits_at_def) have "\ Y \_1 \_2. splits_at ?\ (fst (D ! (N - 1))) \_1 Y \_2" using Derives1_\ splits_at_ex by blast then obtain Y \_1 \_2 where \_split: "splits_at ?\ (fst (D ! (N - 1))) \_1 Y \_2" by blast have NFix: "LeftDerivationFix (\_1 @ [X] @ \_2) (length \_1) (drop N D) j \" using Minimal_N \_combine i_def by auto from LeftDerives1_\_hyp \_split have "\u v. \_1 = \_1 @ u \ \_2 = v @ \_2 \ snd (snd (D ! (N - 1))) = u @ [X] @ v" proof (induct rule: LeftDerives1_X_is_part_of_rule) case (Suffix suffix) let ?k = "N - 1" let ?\ = "Derive a (take ?k D)" let ?i = "length \_1" have k_less: "?k < length D" using "2.hyps" Minimal_N by linarith then have k_leq: "?k \ length D" by auto have drop_k_d: "drop ?k D = (D ! (N - 1))#(drop N D)" using "2.hyps" Cons_nth_drop_Suc k_less by fastforce from LeftDerivationFix_grow_suffix[OF NFix Suffix(4) Suffix(3)] Suffix(1) Suffix(2) 2 have "LeftDerivationFix ?\ ?i (drop ?k D) j \" apply auto by (metis One_nat_def drop_k_d) with Minimal(2)[where k="?k"] show "False" using "2.hyps" k_leq by auto next case (Prefix prefix) have collapse: "(fst (D ! (N - 1)), snd (D ! (N - 1))) # drop N D = drop (N - 1) D" by (metis "2.hyps" Cons_nth_drop_Suc Minimal_N Suc_diff_1 neq0_conv not_less not_less_eq prod.collapse) from LeftDerivationFix_grow_prefix[OF NFix Prefix(2)] Prefix(1) collapse have "LeftDerivationFix ?\ (length prefix) (drop (N - 1) D) j \" by auto with Minimal(2)[where k = "N - 1"] show "False" by (metis Minimal_N collapse diff_le_self le_neq_implies_less less_imp_diff_less less_or_eq_imp_le not_Cons_self2) qed then obtain u v where uv: "\_1 = \_1 @ u \ \_2 = v @ \_2 \ snd (snd (D ! (N - 1))) = u @ [X] @ v" by blast have X_1: "snd (snd (D ! (N - Suc 0))) ! length u = X" using uv by auto have X_2: "\ ! j = X" using LeftDerivationFix_def NFix by auto show ?case apply (rule disjI2) apply (rule_tac x="N - 1" in exI) apply (rule_tac x="?\" in exI) apply (rule_tac x="length u" in exI) apply (rule conjI) using Minimal_N less2 apply linarith apply (rule conjI) using LeftDerivation_take_derive less.prems(1) apply blast apply (subst LeftDerivationIntro_def) apply (rule_tac x=\ in exI) apply auto using LeftDerives1_\ One_nat_def apply presburger using uv apply auto[1] using X_1 X_2 apply auto[1] by (metis (no_types, lifting) "2.hyps" Derives1_\ Derives1_split Minimal_N One_nat_def Suc_diff_1 \_split append_eq_conv_conj i_def length_append neq0_conv splits_at_def uv) qed qed qed qed type_synonym deriv = "nat \ nat \ nat" type_synonym ladder = "deriv list" definition deriv_n :: "deriv \ nat" where "deriv_n d = fst d" definition deriv_j :: "deriv \ nat" where "deriv_j d = fst (snd d)" definition deriv_ix :: "deriv \ nat" where "deriv_ix d = snd (snd d)" definition deriv_i :: "deriv \ nat" where "deriv_i d = snd (snd d)" definition ladder_j :: "ladder \ nat \ nat" where "ladder_j L index = deriv_j (L ! index)" definition ladder_i :: "ladder \ nat \ nat" where "ladder_i L index = (if index = 0 then deriv_i (hd L) else ladder_j L (index - 1))" definition ladder_n :: "ladder \ nat \ nat" where "ladder_n L index = deriv_n (L ! index)" definition ladder_prev_n :: "ladder \ nat \ nat" where "ladder_prev_n L index = (if index = 0 then 0 else (ladder_n L (index - 1)))" definition ladder_ix :: "ladder \ nat \ nat" where "ladder_ix L index = (if index = 0 then undefined else deriv_ix (L ! index))" definition ladder_last_j :: "ladder \ nat" where "ladder_last_j L = ladder_j L (length L - 1)" definition ladder_last_n :: "ladder \ nat" where "ladder_last_n L = ladder_n L (length L - 1)" definition is_ladder :: "derivation \ ladder \ bool" where "is_ladder D L = (L \ [] \ (\ u. u < length L \ ladder_n L u \ length D) \ (\ u v. u < v \ v < length L \ ladder_n L u < ladder_n L v) \ ladder_last_n L = length D)" definition ladder_\ :: "sentence \ derivation \ ladder \ nat \ sentence" where "ladder_\ a D L index = Derive a (take (ladder_n L index) D)" definition ladder_\ :: "sentence \ derivation \ ladder \ nat \ sentence" where "ladder_\ a D L index = (if index = 0 then a else ladder_\ a D L (index - 1))" definition LeftDerivationIntrosAt :: "sentence \ derivation \ ladder \ nat \ bool" where "LeftDerivationIntrosAt a D L index = ( let \ = ladder_\ a D L index in let i = ladder_i L index in let j = ladder_j L index in let ix = ladder_ix L index in let \ = ladder_\ a D L index in let n = ladder_n L (index - 1) in let m = ladder_n L index in let e = D ! n in let E = drop (Suc n) (take m D) in i = fst e \ LeftDerivationIntro \ i (snd e) ix E j \)" definition LeftDerivationIntros :: "sentence \ derivation \ ladder \ bool" where "LeftDerivationIntros a D L = ( \ index. 1 \ index \ index < length L \ LeftDerivationIntrosAt a D L index)" definition LeftDerivationLadder :: "sentence \ derivation \ ladder \ sentence \ bool" where "LeftDerivationLadder a D L b = ( LeftDerivation a D b \ is_ladder D L \ LeftDerivationFix a (ladder_i L 0) (take (ladder_n L 0) D) (ladder_j L 0) (ladder_\ a D L 0) \ LeftDerivationIntros a D L)" definition mk_deriv_fix :: "nat \ nat \ nat \ deriv" where "mk_deriv_fix i n j = (n, j, i)" definition mk_deriv_intro :: "nat \ nat \ nat \ deriv" where "mk_deriv_intro ix n j = (n, j, ix)" lemma mk_deriv_fix_i[simp]: "deriv_i (mk_deriv_fix i n j) = i" by (simp add: deriv_i_def mk_deriv_fix_def) lemma mk_deriv_fix_j[simp]: "deriv_j (mk_deriv_fix i n j) = j" by (simp add: deriv_j_def mk_deriv_fix_def) lemma mk_deriv_fix_n[simp]: "deriv_n (mk_deriv_fix i n j) = n" by (simp add: deriv_n_def mk_deriv_fix_def) lemma mk_deriv_intro_i[simp]: "deriv_i (mk_deriv_intro i n j) = i" by (simp add: deriv_i_def mk_deriv_intro_def) lemma mk_deriv_intro_ix[simp]: "deriv_ix (mk_deriv_intro ix n j) = ix" by (simp add: deriv_ix_def mk_deriv_intro_def) lemma mk_deriv_intro_j[simp]: "deriv_j (mk_deriv_intro i n j) = j" by (simp add: deriv_j_def mk_deriv_intro_def) lemma mk_deriv_intro_n[simp]: "deriv_n (mk_deriv_intro i n j) = n" by (simp add: deriv_n_def mk_deriv_intro_def) lemma LeftDerivationFix_implies_ex_ladder: "LeftDerivationFix a i D j \ \ \ L. LeftDerivationLadder a D L \ \ ladder_last_j L = j \ ladder_last_n L = length D" apply (rule_tac x="[mk_deriv_fix i (length D) j]" in exI) apply (auto simp add: LeftDerivationLadder_def) apply (simp add: LeftDerivationFix_def) apply (simp add: is_ladder_def) apply (auto simp add: ladder_i_def ladder_j_def ladder_n_def ladder_\_def) apply (simp add: ladder_last_n_def ladder_n_def) using Derive LeftDerivationFix_def LeftDerivation_implies_Derivation apply blast apply (simp add: LeftDerivationIntros_def) apply (simp add: ladder_last_j_def ladder_j_def) apply (simp add: ladder_last_n_def ladder_n_def) done lemma trivP[case_names prems]: "P \ P" by blast lemma LeftDerivationLadder_ladder_n_bound: assumes "LeftDerivationLadder a D L b" assumes "index < length L" shows "ladder_n L index \ length D" using LeftDerivationLadder_def assms(1) assms(2) is_ladder_def by blast lemma LeftDerivationLadder_deriv_n_bound: assumes "LeftDerivationLadder a D L b" assumes "index < length L" shows "deriv_n (L ! index) \ length D" using LeftDerivationLadder_def assms(1) assms(2) is_ladder_def ladder_n_def by auto lemma ladder_n_simp1[simp]: "u < length L \ ladder_n (L@L') u = ladder_n L u" by (simp add: ladder_n_def) lemma ladder_n_simp2[simp]: "ladder_n (L@[d]) (length L) = deriv_n d" by (simp add: ladder_n_def) lemma ladder_j_simp1[simp]: "u < length L \ ladder_j (L@L') u = ladder_j L u" by (simp add: ladder_j_def) lemma ladder_j_simp2[simp]: "ladder_j (L@[d]) (length L) = deriv_j d" by (simp add: ladder_j_def) lemma ladder_i_simp1[simp]: "u < length L \ ladder_i (L@L') u = ladder_i L u" by (auto simp add: ladder_i_def) lemma ladder_ix_simp1[simp]: "u < length L \ ladder_ix (L@L') u = ladder_ix L u" by (auto simp add: ladder_ix_def) lemma ladder_ix_simp2[simp]: "L \ [] \ ladder_ix (L@[d]) (length L) = deriv_ix d" by (auto simp add: ladder_ix_def) lemma ladder_\_simp1[simp]: "u < length L \ ladder_\ a D (L@L') u = ladder_\ a D L u" by (simp add: ladder_\_def) lemma ladder_\_simp2[simp]: "u < length L \ is_ladder D L \ ladder_\ a (D@D') L u = ladder_\ a D L u" by (simp add: is_ladder_def ladder_\_def) lemma ladder_\_simp1[simp]: "u < length L \ ladder_\ a D (L@L') u = ladder_\ a D L u" by (simp add: ladder_\_def) lemma ladder_\_simp2[simp]: "u < length L \ is_ladder D L \ ladder_\ a (D@D') L u = ladder_\ a D L u" by (simp add: is_ladder_def ladder_\_def) lemma ladder_n_minus_1_bound: "is_ladder D L \ index \ 1 \ index < length L \ ladder_n L (index - Suc 0) < length D" by (metis (no_types, lifting) One_nat_def Suc_diff_1 Suc_le_lessD dual_order.strict_implies_order is_ladder_def le_neq_implies_less not_less) lemma LeftDerivationIntrosAt_ignore_appendix: assumes is_ladder: "is_ladder D L" assumes hyp: "LeftDerivationIntrosAt a D L index" assumes index_ge: "index \ 1" assumes index_less: "index < length L" shows "LeftDerivationIntrosAt a (D @ D') (L @ L') index" proof - have index_minus_1: "index - Suc 0 < length L" using index_less by arith have is_0: "ladder_n L index - length D = 0" using index_less is_ladder is_ladder_def by auto from index_ge index_less show ?thesis apply (simp add: LeftDerivationIntrosAt_def Let_def) apply (simp add: index_minus_1 is_ladder ladder_n_minus_1_bound is_0) using hyp apply (auto simp add: LeftDerivationIntrosAt_def Let_def) done qed lemma ladder_i_eq_last_j: "L \ [] \ ladder_i (L @ L') (length L) = ladder_last_j L" by (simp add: ladder_i_def ladder_last_j_def) lemma ladder_last_n_intro: "L \ [] \ ladder_n L (length L - Suc 0) = ladder_last_n L" by (simp add: ladder_last_n_def) lemma is_ladder_not_empty: "is_ladder D L \ L \ []" using is_ladder_def by blast lemma last_ladder_\: assumes is_ladder: "is_ladder D L" assumes ladder_last_n: "ladder_last_n L = length D" shows "ladder_\ a D L (length L - Suc 0) = Derive a D" proof - from is_ladder is_ladder_not_empty have "L \ []" by blast then show ?thesis by (simp add: ladder_\_def ladder_last_n_intro ladder_last_n) qed lemma ladder_\_full: assumes is_ladder: "is_ladder D L" assumes ladder_last_n: "ladder_last_n L = length D" shows "ladder_\ a (D @ D') (L @ L') (length L) = Derive a D" proof - from is_ladder have L_not_empty: "L \ []" by (simp add: is_ladder_def) with is_ladder ladder_last_n show ?thesis apply (simp add: ladder_\_def) apply (simp add: last_ladder_\) done qed lemma LeftDerivationIntro_implies_LeftDerivation: "LeftDerivationIntro \ i r ix D j \ \ LeftDerivation \ ((i,r)#D) \" using LeftDerivationFix_def LeftDerivationIntro_def by auto lemma LeftDerivationLadder_grow: "LeftDerivationLadder a D L \ \ ladder_last_j L = i \ LeftDerivationIntro \ i r ix E j \ \ LeftDerivationLadder a (D@[(i, r)]@E) (L@[mk_deriv_intro ix (Suc(length D + length E)) j]) \" proof (induct arbitrary: a D L \ i r ix E j \ rule: trivP) case prems { fix u :: nat assume "u < Suc (length L)" then have "u < length L \ u = length L" by arith then have "ladder_n (L @ [mk_deriv_intro ix (Suc (length D + length E)) j]) u \ Suc (length D + length E)" proof (induct rule: disjCases2) case 1 then show ?case apply simp by (meson LeftDerivationLadder_ladder_n_bound le_Suc_eq le_add1 le_trans prems(1)) next case 2 then show ?case by (simp add: ladder_n_def) qed } note ladder_n_ineqs = this { fix u :: nat fix v :: nat assume u_less_v: "u < v" assume "v < Suc (length L)" then have "v < length L \ v = length L" by arith then have "ladder_n (L @ [mk_deriv_intro ix (Suc (length D + length E)) j]) u < ladder_n (L @ [mk_deriv_intro ix (Suc (length D + length E)) j]) v" proof (induct rule: disjCases2) case 1 with u_less_v have u_bound: "u < length L" by arith show ?case using 1 u_bound apply simp using prems u_less_v LeftDerivationLadder_def is_ladder_def by auto next case 2 with u_less_v have u_bound: "u < length L" by arith have "deriv_n (L ! u) \ length D" using LeftDerivationLadder_deriv_n_bound prems(1) u_bound by blast then show ?case apply (simp add: u_bound) apply (simp add: ladder_n_def 2) done qed } note ladder_n_ineqs = ladder_n_ineqs this have is_ladder: "is_ladder (D @ (i, r) # E) (L @ [mk_deriv_intro ix (Suc (length D + length E)) j])" apply (auto simp add: is_ladder_def) using ladder_n_ineqs apply auto apply (simp add: ladder_last_n_def) done have is_ladder_L: "is_ladder D L" using LeftDerivationLadder_def prems.prems(1) by blast have ladder_last_n_eq_length: "ladder_last_n L = length D" using is_ladder_L is_ladder_def by blast have L_not_empty: "L \ []" using LeftDerivationLadder_def is_ladder_def prems(1) by blast { fix index :: nat assume index_ge: "Suc 0 \ index" assume "index < Suc (length L)" then have "index < length L \ index = length L" by arith then have "LeftDerivationIntrosAt a (D @ (i, r) # E) (L @ [mk_deriv_intro ix (Suc (length D + length E)) j]) index" proof (induct rule: disjCases2) case 1 then show ?case using LeftDerivationIntrosAt_ignore_appendix LeftDerivationIntros_def LeftDerivationLadder_def One_nat_def index_ge prems.prems(1) by presburger next case 2 have min_simp: "\ n E. min n (Suc (n + length E)) = n" by auto with 2 prems is_ladder_L ladder_last_n_eq_length show ?case apply (simp add: LeftDerivationIntrosAt_def) apply (simp add: L_not_empty ladder_i_eq_last_j ladder_last_n_intro) apply (simp add: ladder_\_full min_simp) apply (simp add: ladder_\_def) by (metis Derive LeftDerivationIntro_implies_LeftDerivation LeftDerivationLadder_def LeftDerivation_implies_Derivation LeftDerivation_implies_append) qed } then show ?case apply (auto simp add: LeftDerivationLadder_def) using prems apply (auto simp add: LeftDerivationLadder_def)[1] using LeftDerivationFix_def LeftDerivationIntro_def LeftDerivation_append apply auto[1] using is_ladder apply simp using L_not_empty apply simp using LeftDerivationLadder_def LeftDerivationLadder_ladder_n_bound ladder_\_def prems.prems(1) apply auto[1] apply (subst LeftDerivationIntros_def) apply auto done qed lemma LeftDerivationIntro_bounds_ij: "LeftDerivationIntro \ i r ix D j \ \ i < length \ \ j < length \" by (meson Derives1_bound LeftDerivationFix_def LeftDerivationIntro_def LeftDerives1_implies_Derives1) theorem LeftDerivationLadder_exists: "LeftDerivation a D \ \ is_sentence \ \ j < length \ \ \ L. LeftDerivationLadder a D L \ \ ladder_last_j L = j" proof (induct "length D" arbitrary: a D \ j rule: less_induct) case less from LeftDerivationFixOrIntro[OF less(2,3,4)] show ?case proof (induct rule: disjCases2) case 1 then obtain i where "LeftDerivationFix a i D j \" by blast show ?case using "1.hyps" LeftDerivationFix_implies_ex_ladder by blast next case 2 then obtain d \ ix where inductrule: "d < length D \ LeftDerivation a (take d D) \ \ LeftDerivationIntro \ (fst (D ! d)) (snd (D ! d)) ix (drop (Suc d) D) j \" by blast then have less_length_D: "length (take d D) < length D" and LeftDerivation_\: "LeftDerivation a (take d D) \" by auto have is_sentence_\: "is_sentence \" using LeftDerivationIntro_is_sentence inductrule by blast have "fst (D ! d) < length \" using LeftDerivationIntro_bounds_ij inductrule by blast from less(1)[OF less_length_D LeftDerivation_\ is_sentence_\, where j=" fst (D ! d)", OF this] obtain L where induct_Ladder: "LeftDerivationLadder a (take d D) L \" and induct_last: "ladder_last_j L = fst (D ! d)" by blast have induct_intro: "LeftDerivationIntro \ (fst (D ! d)) (snd (D ! d)) ix (drop (Suc d) D) j \" using inductrule by blast have "d < length D" using inductrule by blast then have simp_to_D: "take d D @ D ! d # drop (Suc d) D = D" using id_take_nth_drop by force from LeftDerivationLadder_grow[OF induct_Ladder induct_last induct_intro] simp_to_D show ?case apply auto apply (rule_tac x= "L @ [mk_deriv_intro ix (Suc (min (length D) d + (length D - Suc d))) j]" in exI) apply (simp add: ladder_last_j_def) done qed qed lemma LeftDerivationLadder_L_0: assumes "LeftDerivationLadder \ D L \" assumes "length L = 1" shows "\ i. LeftDerivationFix \ i D (ladder_last_j L) \" proof - have "is_ladder D L" using assms by (auto simp add: LeftDerivationLadder_def) then have ladder_n: "ladder_n L 0 = length D" by (simp add: assms(2) is_ladder_def ladder_last_n_def) show ?thesis apply (rule_tac x = "ladder_i L 0" in exI) using assms(1) apply (auto simp add: LeftDerivationLadder_def) by (metis Derive LeftDerivationFix_def LeftDerivation_implies_Derivation One_nat_def assms(2) diff_Suc_1 ladder_last_j_def ladder_n order_refl take_all) qed lemma LeftDerivationFix_splits_at_derives: assumes "LeftDerivationFix a i D j b" shows "\ U a1 a2 b1 b2. splits_at a i a1 U a2 \ splits_at b j b1 U b2 \ derives a1 b1 \ derives a2 b2" proof - note hyp = LeftDerivationFix_def[where \=a and \=b and D=D and i=i and j=j] from hyp obtain E F where EF: "D = E @ derivation_shift F 0 (Suc j) \ LeftDerivation (take i a) E (take j b) \ LeftDerivation (drop (Suc i) a) F (drop (Suc j) b)" using assms by blast show ?thesis apply (rule_tac x="a ! i" in exI) apply (rule_tac x="take i a" in exI) apply (rule_tac x="drop (Suc i) a" in exI) apply (rule_tac x="take j b" in exI) apply (rule_tac x="drop (Suc j) b" in exI) using Derivation_implies_derives LeftDerivation_implies_Derivation assms hyp splits_at_def by blast qed lemma LeftDerivation_append_suffix: "LeftDerivation a D b \ is_sentence c \ LeftDerivation (a@c) D (b@c)" proof (induct D arbitrary: a b c) case Nil then show ?case by auto next case (Cons d D) then show ?case apply auto apply (rule_tac x="x@c" in exI) apply auto using LeftDerives1_append_suffix by simp qed lemma LeftDerivation_impossible: "LeftDerivation a D b \ i < length a \ is_nonterminal (a ! i) \ derivation_ge D (Suc i) \ D = []" proof (induct D) case Nil then show ?case by auto next case (Cons d D) then have lm: "\ j. leftmost j a \ j \ i" by (metis Derives1_sentence1 LeftDerivation.simps(2) LeftDerives1_implies_Derives1 leftmost_exists leftmost_unique) from Cons show ?case apply auto apply (auto simp add: derivation_ge_def LeftDerives1_def) using lm[where j="fst d"] by arith qed lemma derivation_ge_shift: "derivation_ge (derivation_shift F 0 j) j" apply (induct F) apply (auto simp add: derivation_ge_def) done lemma LeftDerivationFix_splits_at_nonterminal: assumes "LeftDerivationFix a i D j b" assumes "is_nonterminal (a ! i)" shows "\ U a1 a2 b1. splits_at a i a1 U a2 \ splits_at b j b1 U a2 \ LeftDerivation a1 D b1" proof - note hyp = LeftDerivationFix_def[where \=a and \=b and D=D and i=i and j=j] from hyp obtain E F where EF: "D = E @ derivation_shift F 0 (Suc j) \ LeftDerivation (take i a) E (take j b) \ LeftDerivation (drop (Suc i) a) F (drop (Suc j) b)" using assms by blast have "\ \. LeftDerivation a E \ \ LeftDerivation \ (derivation_shift F 0 (Suc j)) b" using EF LeftDerivation_append assms(1) hyp by blast then obtain \ where \_intro: "LeftDerivation a E \ \ LeftDerivation \ (derivation_shift F 0 (Suc j)) b" by blast have "LeftDerivation ((take i a)@(drop i a)) E ((take j b)@(drop i a))" by (metis EF LeftDerivation_append_suffix append_take_drop_id assms(1) hyp is_sentence_concat) then have "LeftDerivation a E ((take j b)@(drop i a))" by simp then have \_decomposed: "\ = (take j b)@(drop i a)" using Derivation_unique_dest LeftDerivation_implies_Derivation \_intro by blast then have "\ ! j = a ! i" by (metis Cons_nth_drop_Suc assms(1) hyp length_take min.absorb2 nth_append_length order.strict_implies_order) then have is_nt: "is_nonterminal (\ ! j)" by (simp add: assms(2)) have index_j: "j < length \" using \_decomposed assms(1) hyp by auto have derivation: "LeftDerivation \ (derivation_shift F 0 (Suc j)) b" by (simp add: \_intro) from LeftDerivation_impossible[OF derivation index_j is_nt derivation_ge_shift] have F: "F = []" by (metis length_0_conv length_derivation_shift) then have \_is_b: "\ = b" using \_intro by auto show ?thesis apply (rule_tac x="a ! i" in exI) apply (rule_tac x="take i a" in exI) apply (rule_tac x="drop (Suc i) a" in exI) apply (rule_tac x="take j b" in exI) using EF F assms(1) hyp splits_at_def by auto qed lemma LeftDerivationIntro_implies_nonterminal: "LeftDerivationIntro \ i (snd e) ix E j \ \ is_nonterminal (\ ! i)" by (simp add: LeftDerivationIntro_def LeftDerives1_def leftmost_is_nonterminal) lemma LeftDerivationIntrosAt_implies_nonterminal: "LeftDerivationIntrosAt a D L index \ is_nonterminal((ladder_\ a D L index) ! (ladder_i L index))" by (meson LeftDerivationIntro_implies_nonterminal LeftDerivationIntrosAt_def) lemma LeftDerivationIntro_examine_rule: "LeftDerivationIntro \ i r ix D j \ \ splits_at \ i \1 M \2 \ \ \. M = fst r \ \ = snd r \ (M, \) \ \" by (metis Derives1_nonterminal Derives1_rule LeftDerivationIntro_def LeftDerives1_implies_Derives1 prod.collapse) lemma LeftDerivation_skip_prefixword_ex: assumes "LeftDerivation (u@v) D w" assumes "is_word u" shows "\ w'. w = u@w' \ LeftDerivation v (derivation_shift D (length u) 0) w'" by (metis LeftDerivation.simps(1) LeftDerivation_breakdown LeftDerivation_implies_Derivation LeftDerivation_skip_prefix append_eq_conv_conj assms(1) assms(2) is_word_Derivation is_word_Derivation_derivation_ge) definition ladder_cut :: "ladder \ nat \ ladder" where "ladder_cut L n = (let i = length L - 1 in L[i := (n, snd (L ! i))])" fun deriv_shift :: "nat \ nat \ deriv \ deriv" where "deriv_shift dn dj (n, j, i) = (n - dn, j - dj, i)" definition ladder_shift :: "ladder \ nat \ nat \ ladder" where "ladder_shift L dn dj = map (deriv_shift dn dj) L" lemma splits_at_append_suffix_prevails: assumes "splits_at (a@b) i u N v" assumes "i < length a" shows "\ v'. v = v'@b \ a=u@[N]@v'" proof - have "min (length a) (Suc i) = Suc i" using Suc_leI assms(2) min.absorb2 by blast then show ?thesis by (metis (no_types) append_assoc append_eq_conv_conj append_take_drop_id assms(1) hd_drop_conv_nth length_take splits_at_def take_hd_drop) qed lemma derivation_shift_right_left_cancel: "derivation_shift (derivation_shift D 0 r) r 0 = D" by (induct D, auto) lemma derivation_shift_left_right_cancel: assumes "derivation_ge D r" shows "derivation_shift (derivation_shift D r 0) 0 r = D" using assms derivation_ge_shift_simp derivation_shift_0_shift by auto lemma LeftDerivation_ge_take: assumes "derivation_ge D k" assumes "LeftDerivation a D b" assumes "D \ []" shows "take k a = take k b \ is_word (take k a)" proof - obtain d D' where d: "d#D' = D" using assms(3) list.exhaust by blast then have "\ x. LeftDerives1 a (fst d) (snd d) x \ LeftDerivation x D' b" using LeftDerivation.simps(2) assms(2) by blast then obtain x where x: "LeftDerives1 a (fst d) (snd d) x \ LeftDerivation x D' b" by blast have fst_d_k: "fst d \ k" using d assms(1) derivation_ge_cons by blast from x fst_d_k have is_word: "is_word (take k a)" by (metis LeftDerives1_def append_take_drop_id is_word_append leftmost_def min.absorb2 take_append take_take) have is_eq: "take k a = take k b" using Derivation_take LeftDerivation_implies_Derivation assms(1) assms(2) by blast show ?thesis using is_word is_eq by blast qed lemma LeftDerivationFix_splits_at_symbol: assumes "LeftDerivationFix a i D j b" shows "\ U a1 a2 b1 b2 n. splits_at a i a1 U a2 \ splits_at b j b1 U b2 \ n \ length D \ LeftDerivation a1 (take n D) b1 \ derivation_ge (drop n D) (Suc(length b1)) \ LeftDerivation a2 (derivation_shift (drop n D) (Suc(length b1)) 0) b2 \ (n = length D \ (n < length D \ is_word (b1@[U])))" proof - note hyp = LeftDerivationFix_def[where \=a and \=b and D=D and i=i and j=j] from hyp obtain E F where EF: "D = E @ derivation_shift F 0 (Suc j) \ LeftDerivation (take i a) E (take j b) \ LeftDerivation (drop (Suc i) a) F (drop (Suc j) b)" using assms by blast have "\ \. LeftDerivation a E \ \ LeftDerivation \ (derivation_shift F 0 (Suc j)) b" using EF LeftDerivation_append assms(1) hyp by blast then obtain \ where \_intro: "LeftDerivation a E \ \ LeftDerivation \ (derivation_shift F 0 (Suc j)) b" by blast have "LeftDerivation ((take i a)@(drop i a)) E ((take j b)@(drop i a))" by (metis EF LeftDerivation_append_suffix append_take_drop_id assms(1) hyp is_sentence_concat) then have "LeftDerivation a E ((take j b)@(drop i a))" by simp then have \_decomposed: "\ = (take j b)@(drop i a)" using Derivation_unique_dest LeftDerivation_implies_Derivation \_intro by blast have derivation: "LeftDerivation \ (derivation_shift F 0 (Suc j)) b" by (simp add: \_intro) have "\ n. n \ length D \ E = take n D" by (metis EF append_eq_conv_conj is_prefix_length is_prefix_take) then obtain n where n: "n \ length D \ E = take n D" by blast have F_def: "drop n D = derivation_shift F 0 (Suc j)" by (metis EF append_eq_conv_conj length_take min.absorb2 n) have min_j: "min (length b) j = j" using assms hyp by linarith have derivation_ge_Suc_j: "derivation_ge (drop n D) (Suc j)" using F_def derivation_ge_shift by simp have LeftDerivation_\_b: "LeftDerivation \ (drop n D) b" by (simp add: F_def \_intro) have is_word_Suc_j_b: "n \ length D \ is_word (take (Suc j) b)" by (metis EF F_def LeftDerivation_ge_take \_intro append_Nil2 derivation_ge_Suc_j length_take min.absorb2 n) have take_Suc_j_b_decompose: "take (Suc j) b = take j b @ [a ! i]" using assms hyp take_Suc_conv_app_nth by auto show ?thesis apply (rule_tac x="a ! i" in exI) apply (rule_tac x="take i a" in exI) apply (rule_tac x="drop (Suc i) a" in exI) apply (rule_tac x="take j b" in exI) apply (rule_tac x="drop (Suc j) b" in exI) apply (rule_tac x="n" in exI) apply (auto simp add: min_j) using assms hyp splits_at_def apply blast using assms hyp splits_at_def apply blast using n apply blast using EF n apply simp using F_def apply simp using derivation_ge_shift apply blast using F_def derivation_shift_right_left_cancel apply simp using EF apply blast using n apply arith using is_word_Suc_j_b take_Suc_j_b_decompose is_word_append apply simp+ done qed lemma LeftDerivation_breakdown': "LeftDerivation (u @ v) D w \ \n w1 w2. n \ length D \ w = w1 @ w2 \ LeftDerivation u (take n D) w1 \ derivation_ge (drop n D) (length w1) \ LeftDerivation v (derivation_shift (drop n D) (length w1) 0) w2" proof - assume hyp: "LeftDerivation (u @ v) D w" from LeftDerivation_breakdown[OF hyp] obtain n w1 w2 where breakdown: "w = w1 @ w2 \ LeftDerivation u (take n D) w1 \ derivation_ge (drop n D) (length w1) \ LeftDerivation v (derivation_shift (drop n D) (length w1) 0) w2" by blast obtain m where m: "m = min (length D) n" by blast have take_m: "take m D = take n D" using m is_prefix_take take_prefix by fastforce have drop_m: "drop m D = drop n D" by (metis append_eq_conv_conj append_take_drop_id length_take m) have m_bound: "m \ length D" by (simp add: m) show ?thesis apply (rule_tac x="m" in exI) apply (rule_tac x="w1" in exI) apply (rule_tac x="w2" in exI) using breakdown m_bound take_m drop_m by auto qed lemma LeftDerives1_append_replace_in_left: assumes ld1: "LeftDerives1 (\@\) i r \" assumes i_bound: "i < length \" shows "\ \'. \ = \'@\ \ LeftDerives1 \ i r \' \ i + length (snd r) \ length \'" proof - obtain \' where \': "\' = (take i \)@(snd r)@(drop (i+1) \)" by blast have fst_r: "fst r = \ ! i" proof - have "\ss n p ssa. \ LeftDerives1 ss n p ssa \ Derives1 ss n p ssa" using LeftDerives1_implies_Derives1 by blast then have "Derives1 (\ @ \) i r \" using ld1 by blast then show ?thesis using Derives1_nonterminal i_bound splits_at_def by auto qed have "Derives1 \ i r \'" using i_bound ld1 apply (auto simp add: \' Derives1_def) apply (rule_tac x="take i \" in exI) apply (rule_tac x="drop (i+1) \" in exI) apply (rule_tac x="fst r" in exI) apply auto apply (simp add: fst_r) using id_take_nth_drop apply blast using Derives1_sentence1 LeftDerives1_implies_Derives1 is_sentence_concat is_sentence_take apply blast apply (metis Derives1_sentence1 LeftDerives1_implies_Derives1 append_take_drop_id is_sentence_concat) using Derives1_rule LeftDerives1_implies_Derives1 by blast then have leftderives1_\_\': "LeftDerives1 \ i r \'" using LeftDerives1_def i_bound ld1 leftmost_cons_less by auto have i_bound_\': "i + length (snd r) \ length \'" using \' i_bound by (simp add: add_mono_thms_linordered_semiring(2) le_add1 less_or_eq_imp_le min.absorb2) have is_sentence_\: "is_sentence \" using Derives1_sentence1 LeftDerives1_implies_Derives1 is_sentence_concat ld1 by blast then have \: "\ = \'@\" using ld1 leftderives1_\_\' Derives1_append_suffix Derives1_unique_dest LeftDerives1_implies_Derives1 by blast show ?thesis apply (rule_tac x="\'" in exI) using \ i_bound_\' leftderives1_\_\' by blast qed lemma LeftDerivationIntro_propagate: assumes intro: "LeftDerivationIntro (\@\) i r ix D j \" assumes i_\: "i < length \" assumes non: "is_nonterminal (\ ! j)" shows "\ \. LeftDerivation \ ((i,r)#D) \ \ \ = \@\ \ j < length \" proof - from intro LeftDerivationIntro_def[where \="\@\" and i=i and r=r and ix=ix and D=D and j=j and \=\] obtain \ where ld_\: "LeftDerives1 (\ @ \) i r \" and ix: "ix < length (snd r) \ snd r ! ix = \ ! j" and \_fix: "LeftDerivationFix \ (i + ix) D j \" by blast from LeftDerives1_append_replace_in_left[OF ld_\ i_\] obtain \' where \': "\ = \' @ \ \ LeftDerives1 \ i r \' \ i + length (snd r) \ length \'" by blast have i_plus_ix_bound: "i + ix < length \'" using \' ix by linarith have ld_\: "LeftDerivationFix (\' @ \) (i + ix) D j \" using \_fix \' by simp then have non_i_ix: "is_nonterminal ((\' @ \) ! (i + ix))" by (simp add: LeftDerivationFix_def non) from LeftDerivationFix_splits_at_nonterminal[OF ld_\ non_i_ix] obtain U a1 a2 b1 where U: "splits_at (\' @ \) (i + ix) a1 U a2 \ splits_at \ j b1 U a2 \ LeftDerivation a1 D b1" by blast have "\ q. a2 = q@\ \ \' = a1 @ [U] @ q" using splits_at_append_suffix_prevails[OF _ i_plus_ix_bound, where b=\] U by blast then obtain q where q: "a2 = q@\ \ \' = a1 @ [U] @ q" by blast show ?thesis apply (rule_tac x="b1@[U]@q" in exI) apply auto apply (rule_tac x="\'" in exI) apply (metis LeftDerivationFix_def LeftDerivation_append_suffix U \' q append_Cons append_Nil is_sentence_concat ld_\) using U q splits_at_combine apply auto[1] using U splits_at_def by auto qed lemma LeftDerivationIntro_finish: assumes intro: "LeftDerivationIntro (\@\) i r ix D j \" assumes i_\: "i < length \" shows "\ k \ \'. k \ length D \ LeftDerivation \ ((i, r)#(take k D)) \ \ LeftDerivation (\ @ \) ((i, r)#(take k D)) (\ @ \) \ derivation_ge (drop k D) (length \) \ LeftDerivation \ (derivation_shift (drop k D) (length \) 0) \' \ \ = \ @ \' \ j < length \" proof - from intro LeftDerivationIntro_def[where \="\@\" and i=i and r=r and ix=ix and D=D and j=j and \=\] obtain \ where ld_\: "LeftDerives1 (\ @ \) i r \" and ix: "ix < length (snd r) \ snd r ! ix = \ ! j" and \_fix: "LeftDerivationFix \ (i + ix) D j \" by blast from LeftDerives1_append_replace_in_left[OF ld_\ i_\] obtain \' where \': "\ = \' @ \ \ LeftDerives1 \ i r \' \ i + length (snd r) \ length \'" by blast have i_plus_ix_bound: "i + ix < length \'" using \' ix by linarith have ld_\: "LeftDerivationFix (\' @ \) (i + ix) D j \" using \_fix \' by simp from LeftDerivationFix_splits_at_symbol[OF ld_\] obtain U a1 a2 b1 b2 n where U: "splits_at (\' @ \) (i + ix) a1 U a2 \ splits_at \ j b1 U b2 \ n \ length D \ LeftDerivation a1 (take n D) b1 \ derivation_ge (drop n D) (Suc (length b1)) \ LeftDerivation a2 (derivation_shift (drop n D) (Suc (length b1)) 0) b2 \ (n = length D \ n < length D \ is_word (b1 @ [U]))" by blast have n_bound: "n \ length D" using U by blast have "\ q. a2 = q@\ \ \' = a1 @ [U] @ q" using splits_at_append_suffix_prevails[OF _ i_plus_ix_bound, where b=\] U by blast then obtain q where q: "a2 = q@\ \ \' = a1 @ [U] @ q" by blast have j: "j = length b1" using U by (simp add: dual_order.strict_implies_order min.absorb2 splits_at_def) have "n = length D \ n < length D \ is_word (b1 @ [U])" using U by blast then show ?thesis proof (induct rule: disjCases2) case 1 from 1 have drop_n_D: "drop n D = []" by (simp add: U) then have "LeftDerivation a2 [] b2" using U by simp then have a2_eq_b2: "a2 = b2" by simp show ?case apply (rule_tac x="n" in exI) apply (rule_tac x="b1@[U]@q" in exI) apply (rule_tac x="\" in exI) apply auto apply (simp add: 1) apply (rule_tac x="\'" in exI) apply (metis LeftDerivationFix_is_sentence LeftDerivation_append_suffix U \' append_Cons append_Nil is_sentence_concat ld_\ q) apply (rule_tac x="\' @ \" in exI) apply (metis "1.hyps" LeftDerivationFix_def U \' a2_eq_b2 id_take_nth_drop ld_\ ld_\ q splits_at_def take_all) apply (simp add: drop_n_D)+ apply (metis U a2_eq_b2 id_take_nth_drop q splits_at_def) using j apply arith done next case 2 obtain E where E: "E = (derivation_shift (drop n D) (Suc (length b1)) 0)" by blast then have "LeftDerivation (q@\) E b2" using U q by simp from LeftDerivation_breakdown'[OF this] obtain n' w1 w2 where w1w2: "n' \ length E \ b2 = w1 @ w2 \ LeftDerivation q (take n' E) w1 \ derivation_ge (drop n' E) (length w1) \ LeftDerivation \ (derivation_shift (drop n' E) (length w1) 0) w2" by blast have length_E_D: "length E = length D - n" using E n_bound by simp have n_plus_n'_bound: "n + n' \ length D" using length_E_D w1w2 n_bound by arith have take_breakdown: "take (n + n') D = (take n D) @ (take n' (drop n D))" using take_add by blast have q_w1: "LeftDerivation q (take n' E) w1" using w1w2 by blast have isw: "is_word (b1 @ [U])" using 2 by blast have take_n': "take n' (drop n D) = derivation_shift (take n' E) 0 (Suc (length b1))" by (metis E U derivation_shift_left_right_cancel take_derivation_shift) have \'_derives_b1_U_w1: "LeftDerivation \' (take (n + n') D) (b1 @ U # w1)" apply (subst take_breakdown) apply (rule_tac LeftDerivation_implies_append[where b="b1@[U]@q"]) apply (metis LeftDerivationFix_is_sentence LeftDerivation_append_suffix U is_sentence_concat ld_\ q) apply (simp add: take_n') by (rule LeftDerivation_append_prefix[OF q_w1, where u = "b1@[U]", OF isw, simplified]) have dge: "derivation_ge (drop (n + n') D) (Suc (length b1 + length w1))" proof - have "derivation_ge (drop n' (drop n D)) (length b1 + 1 + length w1)" by (metis (no_types) E Suc_eq_plus1 U append_take_drop_id derivation_ge_append derivation_ge_shift_plus drop_derivation_shift w1w2) then show "derivation_ge (drop (n + n') D) (Suc (length b1 + length w1))" by (metis (no_types) Suc_eq_plus1 add.commute drop_drop semiring_normalization_rules(23)) qed show ?case apply (rule_tac x="n+n'" in exI) apply (rule_tac x="b1 @ [U] @ w1" in exI) apply (rule_tac x=w2 in exI) apply auto using n_plus_n'_bound apply simp apply (rule_tac x="\'" in exI) using \' \'_derives_b1_U_w1 apply blast apply (rule_tac x="\' @ \" in exI) apply (metis Cons_eq_appendI LeftDerivationFix_is_sentence LeftDerivation_append_suffix \' \'_derives_b1_U_w1 append_assoc is_sentence_concat ld_\ ld_\) apply (rule dge) apply (metis E Suc_eq_plus1 add.commute derivation_shift_0_shift drop_derivation_shift drop_drop w1w2) using U splits_at_combine w1w2 apply auto[1] by (simp add: j) qed qed lemma LeftDerivationLadder_propagate: "LeftDerivationLadder (\@\) D L \ \ ladder_i L 0 < length \ \ n = ladder_n L index \ index < length L \ if (index + 1 < length L) then (\ \. LeftDerivation \ (take n D) \ \ ladder_\ (\@\) D L index = \@\ \ ladder_j L index < length \) else (\ n' \ \'. (index = 0 \ ladder_prev_n L index < n') \ n' \ n \ LeftDerivation \ (take n' D) \ \ LeftDerivation (\@\) (take n' D) (\@\) \ derivation_ge (drop n' D) (length \) \ LeftDerivation \ (derivation_shift (drop n' D) (length \) 0) \' \ ladder_\ (\@\) D L index = \@\' \ ladder_j L index < length \)" proof (induct index arbitrary: n) case 0 have ldfix: "LeftDerivationFix (\@\) (ladder_i L 0) (take n D) (ladder_j L 0) (ladder_\ (\@\) D L 0)" using "0.prems"(1) "0.prems"(3) LeftDerivationLadder_def by blast from 0 have "1 < length L \ 1 = length L" by arith then show ?case proof (induct rule: disjCases2) case 1 have "LeftDerivationIntrosAt (\@\) D L 1" using "0.prems"(1) "1.hyps" LeftDerivationIntros_def LeftDerivationLadder_def by blast from LeftDerivationIntrosAt_implies_nonterminal[OF this] have "is_nonterminal (ladder_\ (\ @ \) D L 0 ! ladder_j L 0)" by (simp add: ladder_\_def ladder_i_def) with ldfix have "is_nonterminal ((\@\) ! (ladder_i L 0))" by (simp add: LeftDerivationFix_def) from LeftDerivationFix_splits_at_nonterminal[OF ldfix this] obtain U a1 a2 b where thesplit: "splits_at (\ @ \) (ladder_i L 0) a1 U a2 \ splits_at (ladder_\ (\ @ \) D L 0) (ladder_j L 0) b U a2 \ LeftDerivation a1 (take n D) b" by blast have "\ \'. a2 = \' @ \ \ \ = a1 @ [U] @ \'" using thesplit splits_at_append_suffix_prevails using "0.prems"(2) by blast then obtain \' where \': "a2 = \' @ \ \ \ = a1 @ ([U] @ \')" by blast obtain \ where \: "\ = b @ ([U] @ \')" by blast have "is_sentence \" using LeftDerivationFix_is_sentence is_sentence_concat ldfix by blast then have "is_sentence ([U] @ \')" using \' is_sentence_concat by blast with \' thesplit have "LeftDerivation (a1 @ ([U] @ \')) (take n D) (b @ ([U] @ \'))" using LeftDerivation_append_suffix by blast then have \_derives_\: "LeftDerivation \ (take n D) \" using \ \' by blast have \_append_\: "ladder_\ (\ @ \) D L 0 = \@\" by (metis \ \' append_assoc splits_at_combine thesplit) have ladder_j_bound: "ladder_j L 0 < length \" by (metis One_nat_def \ diff_add_inverse dual_order.strict_implies_order leD le_add1 length_Cons length_append length_take list.size(3) min.absorb2 neq0_conv splits_at_def thesplit zero_less_diff zero_less_one) show ?case using 1 apply simp apply (rule_tac x="\" in exI) by (auto simp add: \_derives_\ \_append_\ ladder_j_bound) next case 2 note case_2 = 2 have n_def: "n = length D" by (metis "0.prems"(1) "0.prems"(3) "2.hyps" LeftDerivationLadder_def One_nat_def diff_Suc_1 is_ladder_def ladder_last_n_intro) then have take_n_D: "take n D = D" by (simp add: eq_imp_le) from LeftDerivationFix_splits_at_symbol[OF ldfix] obtain U a1 a2 b1 b2 m where U: "splits_at (\ @ \) (ladder_i L 0) a1 U a2 \ splits_at (ladder_\ (\ @ \) D L 0) (ladder_j L 0) b1 U b2 \ m \ length (take n D) \ LeftDerivation a1 (take m (take n D)) b1 \ derivation_ge (drop m (take n D)) (Suc (length b1)) \ LeftDerivation a2 (derivation_shift (drop m (take n D)) (Suc (length b1)) 0) b2 \ (m = length (take n D) \ (m < length (take n D) \ is_word (b1 @ [U])))" by blast obtain D' where D': "D' = derivation_shift (drop m D) (Suc (length b1)) 0" by blast then have a2_derives_b2: "LeftDerivation a2 D' b2" using U take_n_D by auto from U have m_leq_n: "m \ n" by (simp add: "0.prems"(1) "0.prems"(3) "0.prems"(4) LeftDerivationLadder_def is_ladder_def min.absorb2) from U have "splits_at (\ @ \) (ladder_i L 0) a1 U a2" by blast from splits_at_append_suffix_prevails[OF this 0(2)] obtain v' where v': "a2 = v' @ \ \ \ = a1 @ [U] @ v'" by blast have a1_derives_b1: "LeftDerivation a1 (take m D) b1" using m_leq_n U by (metis "0.prems"(1) "0.prems"(3) "2.hyps" LeftDerivationLadder_def One_nat_def cancel_comm_monoid_add_class.diff_cancel is_ladder_def ladder_last_n_intro order_refl take_all) have "LeftDerivation (v' @ \) D' b2" using a2_derives_b2 v' by simp from LeftDerivation_breakdown'[OF this] obtain m' w1 w2 where w12: "b2 = w1 @ w2 \ m' \ length D' \ LeftDerivation v' (take m' D') w1 \ derivation_ge (drop m' D') (length w1) \ LeftDerivation \ (derivation_shift (drop m' D') (length w1) 0) w2" by blast have "length D' \ length D - m" by (simp add: D') then have "m' \ length D - m" using w12 dual_order.trans by blast then have m_m'_leq_n: "m + m' \ n" using n_def m_leq_n le_diff_conv2 add.commute by linarith obtain \ where \: "\ = b1 @ ([U] @ w1)" by blast have "is_sentence ([U] @ v')" using LeftDerivationFix_is_sentence is_sentence_concat ldfix v' by blast then have "LeftDerivation (a1 @ ([U] @ v')) (take m D) (b1 @ ([U] @ v'))" using LeftDerivation_append_suffix a1_derives_b1 by blast then have \_derives_pre_\: "LeftDerivation \ (take m D) (b1 @ ([U] @ v'))" using v' by blast have "m = n \ (m < n \ is_word (b1 @ [U]))" using U n_def[symmetric] take_n_D by simp then have pre_\_derives_\: "LeftDerivation (b1 @ ([U] @ v')) (take m' (drop m D)) \" proof (induct rule: disjCases2) case 1 then have "m' = 0" using m_m'_leq_n by arith then show ?case apply (simp add: \) using w12 by simp next case 2 then have is_word_prefix: "is_word (b1 @ [U])" by blast have take_drop_eq: "take m' (drop m D) = derivation_shift (take m' D') 0 (length (b1 @ [U]))" apply (simp add: D' take_derivation_shift) by (metis U derivation_shift_left_right_cancel take_derivation_shift take_n_D) have v'_derives_w1: "LeftDerivation v' (take m' D') w1" by (simp add: w12) with is_word_prefix have "LeftDerivation ((b1 @ [U]) @ v') (derivation_shift (take m' D') 0 (length (b1 @ [U]))) ((b1 @ [U]) @ w1)" using LeftDerivation_append_prefix by blast with take_drop_eq show ?case by (simp add: \) qed have "(take m D) @ (take m' (drop m D)) = (take (m + m') D)" by (simp add: take_add) then have \_derives_\: "LeftDerivation \ (take (m + m') D) \" using LeftDerivation_implies_append \_derives_pre_\ pre_\_derives_\ by fastforce have derivation_ge_drop_m_m': "derivation_ge (drop (m + m') D) (length \)" proof - have f1: "drop m' (drop m D) = drop (m + m') D" by (simp add: add.commute) have "derivation_ge (drop m' (drop m D)) (Suc (length b1))" by (metis (no_types) U append_take_drop_id derivation_ge_append take_n_D) then show "derivation_ge (drop (m + m') D) (length \)" using f1 by (metis (no_types) D' \ append_assoc derivation_ge_shift_plus drop_derivation_shift length_append length_append_singleton w12) qed have \_derives_w2: "LeftDerivation \ (derivation_shift (drop (m + m') D) (length \) 0) w2" proof - have "derivation_shift (drop m' D') (length w1) 0 = derivation_shift (drop (m + m') D) (length \) 0" by (simp add: D' \ add.commute derivation_shift_0_shift drop_derivation_shift) then show "LeftDerivation \ (derivation_shift (drop (m + m') D) (length \) 0) w2" using w12 by presburger qed have ladder_\_def: "ladder_\ (\ @ \) D L 0 = \ @ w2" using U \ splits_at_combine w12 by auto have ladder_j_bound: "ladder_j L 0 < length \" using U \ splits_at_def by auto show ?case using 2 apply simp apply (rule_tac x="m + m'" in exI) apply (auto simp add: m_m'_leq_n) apply (rule_tac x="\" in exI) apply (auto simp add: \_derives_\) using LeftDerivationFix_is_sentence LeftDerivation_append_suffix \_derives_\ is_sentence_concat ldfix apply blast using derivation_ge_drop_m_m' apply blast apply (rule_tac x="w2" in exI) apply auto using \_derives_w2 apply blast using ladder_\_def apply blast using ladder_j_bound apply blast done qed next case (Suc index) have step: "LeftDerivationIntrosAt (\@\) D L (Suc index)" by (metis LeftDerivationIntros_def LeftDerivationLadder_def Suc.prems(1) Suc.prems(4) Suc_eq_plus1_left le_add1) have index_plus_1_bound: "index + 1 < length L" using Suc.prems(4) by linarith then have index_bound: "index < length L" by arith obtain n' where n': "n' = ladder_n L index" by blast from Suc.hyps[OF Suc.prems(1) Suc.prems(2) n' index_bound] index_plus_1_bound have "\ \'. LeftDerivation \ (take n' D) \' \ ladder_\ (\@\) D L index = \'@\ \ ladder_j L index < length \'" by auto then obtain \' where \': "LeftDerivation \ (take n' D) \' \ ladder_\ (\@\) D L index = \'@\ \ ladder_j L index < length \'" by blast have Suc_index_bound: "Suc index < length L" using Suc.prems by auto have is_ladder: "is_ladder D L" using Suc.prems LeftDerivationLadder_def by auto have n_def: "n = ladder_n L (Suc index)" using Suc_index_bound n' by (simp add: Suc.prems(3)) with n' have n'_less_n: "n' < n" using is_ladder Suc_index_bound is_ladder_def lessI by blast have ladder_\_is_\: "ladder_\ (\@\) D L (Suc index) = ladder_\ (\@\) D L index" by (simp add: ladder_\_def) obtain i where i: "i = ladder_i L (Suc index)" by blast obtain e where e: "e = (D ! n')" by blast obtain E where E: "E = drop (Suc n') (take n D)" by blast obtain ix where ix: "ix = ladder_ix L (Suc index)" by blast obtain j where j: "j = ladder_j L (Suc index)" by blast obtain \ where \: "\ = ladder_\ (\@\) D L (Suc index)" by blast have intro: "LeftDerivationIntro (\'@\) i (snd e) ix E j \" by (metis E LeftDerivationIntrosAt_def \' \ ladder_\_is_\ diff_Suc_1 e i ix j local.step n' n_def) have is_eq_fst_e: "i = fst e" by (metis LeftDerivationIntrosAt_def diff_Suc_1 e i local.step n') have i_less_\': "i < length \'" using i \' ladder_i_def by simp have "(Suc index) + 1 < length L \ (Suc index) + 1 = length L" using Suc_index_bound by arith then show ?case proof (induct rule: disjCases2) case 1 from 1 have "LeftDerivationIntrosAt (\@\) D L (Suc (Suc index))" by (metis LeftDerivationIntros_def LeftDerivationLadder_def Suc.prems(1) Suc_eq_plus1 Suc_eq_plus1_left le_add1) from LeftDerivationIntrosAt_implies_nonterminal[OF this] have "is_nonterminal (ladder_\ (\ @ \) D L (Suc (Suc index)) ! ladder_i L (Suc (Suc index)))" by blast then have non_\_j: "is_nonterminal (\ ! j)" by (simp add: \ j ladder_\_def ladder_i_def) from LeftDerivationIntro_propagate[OF intro i_less_\' non_\_j] obtain \ where \: "LeftDerivation \' ((i, snd e) # E) \ \ \ = \ @ \ \ j < length \" by blast have \_\: "LeftDerivation \ ((take n' D)@((i, snd e) # E)) \" using \' \ LeftDerivation_implies_append by blast have i_e: "(i, snd e) = e" by (simp add: is_eq_fst_e) have take_n_D_e: "((take n' D)@(e # E)) = take n D" - proof - (* automatically found *) - obtain nn :: "(nat \ symbol \ symbol list) list \ (nat \ nat \ nat) list \ nat" and - nna :: "(nat \ symbol \ symbol list) list \ (nat \ nat \ nat) list \ nat" and - nnb :: "(nat \ symbol \ symbol list) list \ (nat \ nat \ nat) list \ nat" where - f1: "(\ps psa. \ is_ladder ps psa \ psa \ [] \ (\n. \ n < length psa \ - ladder_n psa n \ length ps) \ (\n na. (\ n < na \ \ na < length psa) \ - ladder_n psa n < ladder_n psa na) \ ladder_last_n psa = length ps) \ - (\ps psa. (psa = [] \ nn ps psa < length psa \ \ ladder_n psa (nn ps psa) \ - length ps \ (nna ps psa < nnb ps psa \ nnb ps psa < length psa) \ - \ ladder_n psa (nna ps psa) < ladder_n psa (nnb ps psa) \ - ladder_last_n psa \ length ps) \ is_ladder ps psa)" - using is_ladder_def by moura - then have f2: "ladder_last_n L = length D" - using is_ladder by blast + proof - (* automatically found, then modified *) + have f2: "ladder_last_n L = length D" + using is_ladder is_ladder_def by blast have f3: "min (ladder_last_n L) n = n" - using f1 by (metis (no_types) Suc_eq_plus1 index_plus_1_bound is_ladder - min.absorb2 n_def) + using is_ladder_def + by (metis (no_types) Suc_eq_plus1 index_plus_1_bound is_ladder min.absorb2 n_def) then have "take n' (take n D) @ take n D ! n' # E = take n D" using f2 by (metis E id_take_nth_drop length_take n'_less_n) then show ?thesis using f3 f2 by (metis (no_types) append_assoc append_eq_conv_conj dual_order.strict_implies_order e length_take min.absorb2 n'_less_n nth_append) qed from 1 show ?case apply auto apply (rule_tac x=\ in exI) apply auto using \_\ i_e take_n_D_e apply auto[1] using \ \ apply blast using \ j by blast next case 2 from LeftDerivationIntro_finish[OF intro i_less_\'] obtain k \ \' where kw\': "k \ length E \ LeftDerivation \' ((i, snd e) # take k E) \ \ LeftDerivation (\' @ \) ((i, snd e) # take k E) (\ @ \) \ derivation_ge (drop k E) (length \) \ LeftDerivation \ (derivation_shift (drop k E) (length \) 0) \' \ \ = \ @ \' \ j < length \" by blast have ladder_last_n_1: "ladder_last_n L = n" by (metis "2.hyps" Suc_eq_plus1 diff_Suc_1 ladder_last_n_def n_def) from is_ladder have ladder_last_n_2: "ladder_last_n L = length D" using is_ladder_def by blast from ladder_last_n_1 ladder_last_n_2 have n_eq_length_D: "n = length D" by blast have take_split: "take (Suc (n' + k)) D = (take n' D) @ ((i, snd e) # take k E)" apply (simp add: E n_eq_length_D) by (metis (no_types, lifting) Cons_eq_appendI add_Suc append_eq_appendI e is_eq_fst_e n'_less_n n_eq_length_D prod.collapse self_append_conv2 take_Suc_conv_app_nth take_add) have \_\: "LeftDerivation \ (take (Suc (n' + k)) D) \" apply (subst take_split) apply (rule LeftDerivation_implies_append[where b="\'"]) apply (simp add: \') using kw\' by blast have Suc_n'_k_bound: "Suc (n' + k) \ n" using E kw\' n'_less_n by auto[1] from 2 show ?case apply auto apply (rule_tac x="Suc (n' + k)" in exI) apply auto apply (simp add: ladder_prev_n_def n') using Suc_n'_k_bound apply blast apply (rule_tac x="\" in exI) apply auto using \_\ apply blast using \_\ LeftDerivationFix_def LeftDerivationLadder_def LeftDerivation_append_suffix Suc.prems(1) is_sentence_concat apply auto[1] apply (metis E add.commute add_Suc_right drop_drop kw\' n_eq_length_D nat_le_linear take_all) apply (rule_tac x="\'" in exI) apply auto apply (metis E LeftDerivationLadder_ladder_n_bound Suc.prems(1) Suc_index_bound add.commute add_Suc_right drop_drop kw\' n_def n_eq_length_D take_all) using \ kw\' apply blast using j kw\' by blast qed qed lemma ladder_i_of_cut_at_0: assumes L_non_empty: "L \ []" shows "ladder_i (ladder_cut L n) 0 = ladder_i L 0" proof - have "length L \ 0" using L_non_empty by auto then have "length L = 1 \ length L > 1" by arith then show ?thesis proof (induct rule: disjCases2) case 1 then show ?case apply (simp add: ladder_cut_def ladder_i_def deriv_i_def) by (simp add: assms hd_conv_nth) next case 2 then show ?case apply (simp add: ladder_cut_def ladder_i_def deriv_i_def) by (metis diff_is_0_eq hd_conv_nth leD list_update_nonempty nth_list_update_neq) qed qed lemma ladder_last_j_of_cut: assumes L_non_empty: "L \ []" shows "ladder_last_j (ladder_cut L n) = ladder_last_j L" proof - have length_L_nonzero: "length L \ 0" using L_non_empty by auto then have length_ladder_cut: "length (ladder_cut L n) = length L" by (metis ladder_cut_def length_list_update) show ?thesis apply (simp add: ladder_last_j_def length_ladder_cut) apply (simp add: ladder_cut_def ladder_j_def deriv_j_def) by (metis length_L_nonzero diff_less neq0_conv nth_list_update_eq snd_conv zero_less_Suc) qed lemma length_ladder_cut: assumes L_non_empty: "L \ []" shows "length (ladder_cut L n) = length L" by (metis ladder_cut_def length_list_update) lemma ladder_last_n_of_cut: assumes L_non_empty: "L \ []" shows "ladder_last_n (ladder_cut L n) = n" proof - show ?thesis apply (simp add: ladder_last_n_def length_ladder_cut[OF L_non_empty]) apply (simp add: ladder_n_def ladder_cut_def deriv_n_def) by (metis assms diff_Suc_less fst_conv length_greater_0_conv nth_list_update_eq) qed lemma ladder_n_of_cut: assumes L_non_empty: "L \ []" assumes "index < length L - 1" shows "ladder_n (ladder_cut L n) index = ladder_n L index" by (metis assms(2) ladder_cut_def ladder_n_def nat_neq_iff nth_list_update_neq) lemma ladder_n_prev_bound: assumes ladder: "is_ladder D L" assumes u_bound: "u < length L - 1" shows "ladder_n L u \ ladder_prev_n L (length L - 1)" proof - have "ladder_n L u \ ladder_n L (length L - 2)" proof - have "u < Suc (length L - 2)" using u_bound by linarith then show ?thesis by (metis (no_types) diff_Suc_less is_ladder_def ladder leI length_greater_0_conv not_less_eq numeral_2_eq_2 order.order_iff_strict) qed then show ?thesis by (metis One_nat_def Suc_diff_Suc diff_Suc_1 ladder_prev_n_def neq0_conv not_less0 numeral_2_eq_2 u_bound zero_less_diff) qed lemma ladder_n_last_is_length: assumes "is_ladder D L" shows "ladder_n L (length L - 1) = length D" using assms is_ladder_def ladder_last_n_intro by auto lemma derivation_ge_shift_implies_derivation_ge: assumes dge: "derivation_ge (derivation_shift F 0 j) k" shows "derivation_ge F (k - j)" proof - have "\ i r. (i, r) \ set (derivation_shift F 0 j) \ i \ k" using dge derivation_ge_def by auto { fix i :: nat fix r :: "symbol \ (symbol list)" assume ir: "(i, r) \ set F" then have "(i + j, r) \ set (derivation_shift F 0 j)" proof - have "(i + j, r) \ (\p. (fst p - 0 + j, snd p)) ` set F" by (metis (lifting) ir diff_zero image_eqI prod.collapse prod.inject) then show ?thesis by (simp add: derivation_shift_def) qed then have "i + j \ k" using dge derivation_ge_def by auto then have "i \ k - j" by auto } then show ?thesis using derivation_ge_def by auto qed lemma Derives1_bound': "Derives1 a i r b \ i \ length b" by (metis Derives1_bound Derives1_take append_Nil2 append_take_drop_id drop_eq_Nil dual_order.strict_implies_order length_take min.absorb2 nat_le_linear) lemma LeftDerivation_Derives1_last: assumes "LeftDerivation a D b" assumes "D \ []" shows "Derives1 (Derive a (take (length D - 1) D)) (fst (last D)) (snd (last D)) b" by (metis Derive LeftDerivation_Derive_take_LeftDerives1 LeftDerivation_implies_Derivation LeftDerives1_implies_Derives1 assms(1) assms(2) last_conv_nth le_refl length_0_conv take_all) lemma last_of_prefix_in_set: assumes "n < length E" assumes "D = E@F" shows "last E \ set (drop n D)" proof - have f1: "last (drop n E) = last E" by (simp add: assms(1)) have "drop n E \ []" by (metis (no_types) Cons_nth_drop_Suc assms(1) list.simps(3)) then show ?thesis using f1 by (metis (no_types) append.simps(2) append_butlast_last_id append_eq_conv_conj assms(2) drop_append in_set_dropD insertCI list.set(2)) qed lemma LeftDerivationFix_cut_appendix: assumes ldfix: "LeftDerivationFix (\@\) i D j (\@\')" assumes \_\: "LeftDerivation \ (take n D) \" assumes n_bound: "n \ length D" assumes dge: "derivation_ge (drop n D) (length \)" assumes i_in: "i < length \" assumes j_in: "j < length \" shows "LeftDerivationFix \ i (take n D) j \" proof - from LeftDerivationFix_def[where \="\@\" and i=i and D=D and j=j and \="\@\'"] obtain E F where EF: "is_sentence (\ @ \) \ is_sentence (\ @ \') \ LeftDerivation (\ @ \) D (\ @ \') \ i < length (\ @ \) \ j < length (\ @ \') \ (\ @ \) ! i = (\ @ \') ! j \ D = E @ derivation_shift F 0 (Suc j) \ LeftDerivation (take i (\ @ \)) E (take j (\ @ \')) \ LeftDerivation (drop (Suc i) (\ @ \)) F (drop (Suc j) (\ @ \'))" using ldfix by auto with i_in j_in have take_i_E_take_j: "LeftDerivation (take i \) E (take j \)" by (simp add: less_or_eq_imp_le) obtain m where m: "m = length E" by blast { assume n_less_m: "n < m" then have E_nonempty: "E \ []" using gr_implies_not0 list.size(3) m by auto have last_E_in_set: "last E \ set (drop n D)" using last_of_prefix_in_set n_less_m m EF by blast obtain k r where kr: "(k, r) = last E" by (metis old.prod.exhaust) have k_lower_bound: "k \ length \" using dge last_E_in_set kr by (metis derivation_ge_def fst_conv) have "\ \'. Derives1 \' k r (take j \)" using LeftDerivation_Derives1_last take_i_E_take_j by (metis E_nonempty kr fst_conv snd_conv) then have "k \ j" by (metis Derives1_bound' j_in length_take less_imp_le_nat min.absorb2) then have k_upper_bound: "k < length \" using j_in by arith from k_lower_bound k_upper_bound have "False" by arith } then have m_le_n: "m \ n" by arith have take_i_E_take_j: "LeftDerivation (take i \) E (take j \)" by (simp add: take_i_E_take_j) have "take n D = E @ (take (n - m) (derivation_shift F 0 (Suc j)))" using EF m m_le_n by auto then have take_n_D: "take n D = E @ (derivation_shift (take (n - m) F) 0 (Suc j))" using take_derivation_shift by auto obtain F' where F': "F' = derivation_shift (take (n - m) F) 0 (Suc j)" by blast have "LeftDerivation ((take i \)@(drop i \)) E ((take j \)@(drop i \))" using take_i_E_take_j by (metis EF LeftDerivation_append_suffix append_take_drop_id is_sentence_concat) then have "LeftDerivation \ E ((take j \)@(drop i \))" by simp with take_n_D have take_j_drop_i: "LeftDerivation ((take j \)@(drop i \)) F' \" using F' by (metis Derivation_unique_dest LeftDerivation_append LeftDerivation_implies_Derivation \_\) have F'_ge: "derivation_ge F' (Suc j)" using F' derivation_ge_shift by blast have drop_append: "drop i \ = [\!i] @ (drop (Suc i) \)" by (simp add: Cons_nth_drop_Suc i_in) have take_append: "take j \ @ [\!i] = take (Suc j) \" by (metis LeftDerivationFix_def i_in j_in ldfix nth_superfluous_append take_Suc_conv_app_nth) have take_drop_Suc: "(take j \)@(drop i \) = (take (Suc j) \)@(drop (Suc i) \)" by (simp add: drop_append take_append) with take_drop_Suc take_j_drop_i have "LeftDerivation ((take (Suc j) \)@(drop (Suc i) \)) F' \" by auto note helper = LeftDerivation_skip_prefix[OF this] have len_take: "length (take (Suc j) \) = Suc j" by (simp add: Suc_leI j_in min.absorb2) have F'_shift: "derivation_shift F' (Suc j) 0 = take (n - m) F" using F' derivation_shift_right_left_cancel by blast have LeftDerivation_drop: "LeftDerivation (drop (Suc i) \) (take (n - m) F) (drop (Suc j) \)" using helper len_take F'_shift F'_ge by auto show ?thesis apply (auto simp add: LeftDerivationFix_def) using LeftDerivationFix_is_sentence is_sentence_concat ldfix apply blast using LeftDerivationFix_is_sentence is_sentence_concat ldfix apply blast using \_\ apply blast using i_in apply blast using j_in apply blast using LeftDerivationFix_def i_in j_in ldfix apply auto[1] apply (rule_tac x=E in exI) apply (rule_tac x="take (n - m) F" in exI) apply auto using take_n_D apply blast using take_i_E_take_j apply blast using LeftDerivation_drop by blast qed lemma LeftDerivationFix_cut_appendix': assumes ldfix: "LeftDerivationFix (\@\) i D j (\@\')" assumes \_\: "LeftDerivation \ D \" assumes i_in: "i < length \" assumes j_in: "j < length \" shows "LeftDerivationFix \ i D j \" proof - obtain n where n: "n = length D" by blast have "LeftDerivationFix \ i (take n D) j \" apply (rule_tac LeftDerivationFix_cut_appendix) apply (rule ldfix) using \_\ n apply auto[1] using n apply auto[1] using n apply auto[1] using i_in apply blast using j_in apply blast done then show ?thesis using n by auto qed lemma LeftDerivationIntro_cut_appendix: assumes ldfix: "LeftDerivationIntro (\@\) i r ix D j (\@\')" assumes \_\: "LeftDerivation \ ((i,r)#(take n D)) \" assumes n_bound: "n \ length D" assumes dge: "derivation_ge (drop n D) (length \)" assumes i_in: "i < length \" assumes j_in: "j < length \" shows "LeftDerivationIntro \ i r ix (take n D) j \" proof - from LeftDerivationIntro_def[where \="\@\" and i=i and r=r and ix=ix and D=D and j=j and \="\@\'"] obtain \ where \: "LeftDerives1 (\ @ \) i r \ \ ix < length (snd r) \ snd r ! ix = (\ @ \') ! j \ LeftDerivationFix \ (i + ix) D j (\ @ \')" using ldfix by blast then have "\ \'. \ = \' @ \ \ i + length (snd r) \ length \'" using i_in using LeftDerives1_append_replace_in_left by blast then obtain \' where \': "\ = \' @ \ \ i + length (snd r) \ length \'" by blast have \_\': "LeftDerives1 \ i r \'" using \' \ using LeftDerives1_skip_suffix i_in by blast from \_\ obtain u where u: "LeftDerives1 \ i r u \ LeftDerivation u (take n D) \" by auto with \_\' have "u = \'" using Derives1_unique_dest LeftDerives1_implies_Derives1 by blast with u have \'_\: "LeftDerivation \' (take n D) \" by auto have ldfix_append: "LeftDerivationFix (\' @ \) (i + ix) D j (\ @ \')" using \' \ by blast have i_plus_ix_bound: "i + ix < length \'" using \ \' using add_lessD1 le_add_diff_inverse less_asym' linorder_neqE_nat nat_add_left_cancel_less by linarith from LeftDerivationFix_cut_appendix[OF ldfix_append \'_\ n_bound dge i_plus_ix_bound j_in] have ldfix: "LeftDerivationFix \' (i + ix) (take n D) j \" . show ?thesis apply (simp add: LeftDerivationIntro_def) apply (rule_tac x="\'" in exI) apply auto using \_\' apply blast using \ apply blast apply (simp add: \ j_in) using ldfix by blast qed lemma LeftDerivationIntro_cut_appendix': assumes ldfix: "LeftDerivationIntro (\@\) i r ix D j (\@\')" assumes \_\: "LeftDerivation \ ((i,r)#D) \" assumes i_in: "i < length \" assumes j_in: "j < length \" shows "LeftDerivationIntro \ i r ix D j \" proof - obtain n where n: "n = length D" by blast have "LeftDerivationIntro \ i r ix (take n D) j \" apply (rule_tac LeftDerivationIntro_cut_appendix) apply (rule ldfix) using \_\ n apply auto[1] using n apply auto[1] using n apply auto[1] using i_in apply blast using j_in apply blast done then show ?thesis using n by auto qed lemma ladder_n_monotone: "is_ladder D L \ u \ v \ v < length L \ ladder_n L u \ ladder_n L v" by (metis is_ladder_def le_neq_implies_less linear not_less) lemma ladder_i_cut: assumes index_bound: "index < length L" shows "ladder_i (ladder_cut L n) index = ladder_i L index" proof - have "index = 0 \ index > 0" by arith then show ?thesis proof (induct rule: disjCases2) case 1 with index_bound have "L \ []" by (simp add: less_numeral_extra(3)) then show ?case using 1 by (simp add: ladder_i_of_cut_at_0) next case 2 then show ?case apply (auto simp add: ladder_i_def ladder_cut_def ladder_j_def deriv_j_def Let_def) using index_bound by auto qed qed lemma ladder_j_cut: assumes index_bound: "index < length L" shows "ladder_j (ladder_cut L n) index = ladder_j L index" by (metis gr_implies_not0 index_bound ladder_cut_def ladder_j_def ladder_last_j_def ladder_last_j_of_cut length_ladder_cut list.size(3) nth_list_update_neq) lemma ladder_ix_cut: assumes index_lower_bound: "index > 0" assumes index_upper_bound: "index < length L" shows "ladder_ix (ladder_cut L n) index = ladder_ix L index" proof - show ?thesis using index_lower_bound apply (simp add: ladder_ix_def deriv_ix_def) by (metis index_upper_bound ladder_cut_def nth_list_update_eq nth_list_update_neq snd_conv) qed lemma LeftDerivation_from_in_between: assumes \_\: "LeftDerivation \ (take u D) \" assumes \_\: "LeftDerivation \ (take v D) \" assumes u_le_v: "u \ v" shows "LeftDerivation \ (drop u (take v D)) \" proof - have take_split: "take v D = take u D @ (drop u (take v D))" by (metis u_le_v add_diff_cancel_left' drop_take le_Suc_ex take_add) then show ?thesis using \_\ \_\ by (metis (no_types, lifting) Derivation_unique_dest LeftDerivation_append LeftDerivation_implies_Derivation) qed lemma LeftDerivationLadder_cut_appendix_helper: assumes LDLadder: "LeftDerivationLadder (\@\) D L \" assumes ladder_i_in_\: "ladder_i L 0 < length \" shows "\ E F \1 \2 L'. D = E@F \ \ = \1 @ \2 \ LeftDerivationLadder \ E L' \1 \ derivation_ge F (length \1) \ LeftDerivation \ (derivation_shift F (length \1) 0) \2 \ L' = ladder_cut L (length E)" proof - have is_ladder_D_L: "is_ladder D L" using LDLadder LeftDerivationLadder_def by blast obtain n where n: "n = ladder_last_n L" by blast then have n_eq_ladder_n: "n = ladder_n L (length L - 1)" using ladder_last_n_def by blast have length_L_nonzero: "length L > 0" using LeftDerivationLadder_def assms(1) is_ladder_def by blast from LeftDerivationLadder_propagate[OF LDLadder ladder_i_in_\ n_eq_ladder_n] obtain n' \ \' where finish: "(length L - 1 = 0 \ ladder_prev_n L (length L - 1) < n') \ n' \ n \ LeftDerivation \ (take n' D) \ \ LeftDerivation (\ @ \) (take n' D) (\ @ \) \ derivation_ge (drop n' D) (length \) \ LeftDerivation \ (derivation_shift (drop n' D) (length \) 0) \' \ ladder_\ (\ @ \) D L (length L - 1) = \ @ \' \ ladder_j L (length L - 1) < length \" using length_L_nonzero by auto obtain E where E: "E = take n' D" by blast obtain F where F: "F = drop n' D" by blast obtain L' where L': "L' = ladder_cut L (length E)" by blast have \_ladder: "\ = ladder_\ (\ @ \) D L (length L - 1)" by (metis Derive LDLadder LeftDerivationLadder_def LeftDerivation_implies_Derivation append_Nil2 append_take_drop_id drop_eq_Nil is_ladder_def ladder_\_def le_refl n n_eq_ladder_n) then have \: "\ = \ @ \'" using finish by auto have "is_sentence (\@\)" using LDLadder LeftDerivationFix_is_sentence LeftDerivationLadder_def by blast then have is_sentence_\: "is_sentence \" using is_sentence_concat by blast have "is_sentence \" using Derivation_implies_derives LDLadder LeftDerivationFix_is_sentence LeftDerivationLadder_def LeftDerivation_implies_Derivation derives_is_sentence by blast then have is_sentence_\: "is_sentence \" using \ is_sentence_concat by blast have length_L': "length L' = length L" by (metis L' ladder_cut_def length_list_update) have ladder_last_n_L': "ladder_last_n L' = length E" using L' ladder_last_n_of_cut length_L_nonzero by blast have length_D_eq_n: "length D = n" using LDLadder LeftDerivationLadder_def is_ladder_def n by auto then have length_E_eq_n': "length E = n'" by (simp add: E finish min.absorb2) { fix u :: nat assume "u < length L'" then have "u < length L' - 1 \ u = length L' - 1" by arith then have "ladder_n L' u \ length E" proof (induct rule: disjCases2) case 1 have u_bound: "u < length L - 1" using 1 by (simp add: length_L') then have L'_eq_L: "ladder_n L' u = ladder_n L u" using L' ladder_n_of_cut length_L_nonzero length_greater_0_conv by blast from u_bound have "ladder_n L u \ ladder_prev_n L (length L - 1)" using ladder_n_prev_bound LeftDerivationLadder_def assms(1) by blast then show ?case using L'_eq_L finish length_E_eq_n' u_bound by linarith next case 2 then have "ladder_n L' u = length E" using ladder_last_n_L' ladder_last_n_def by auto then show ?case by auto qed } note ladder_n_bound = this { fix u :: nat fix v :: nat assume u_less_v: "u < v" assume v_bound: "v < length L'" have "v < length L' - 1 \ v = length L' - 1" using v_bound by arith then have "ladder_n L' u < ladder_n L' v" proof (induct rule: disjCases2) case 1 show ?case using "1.hyps" L' LeftDerivationLadder_def assms(1) is_ladder_def ladder_n_of_cut length_L' u_less_v by auto next case 2 note v_def = 2 have "v = 0 \ v \ 0" by arith then show ?case proof (induct rule: disjCases2) case 1 then show ?case using u_less_v by auto next case 2 then have "ladder_prev_n L (length L - 1) < n'" using finish v_def length_L' by linarith then show ?case by (metis (no_types, lifting) L' LeftDerivationLadder_def assms(1) ladder_last_n_L' ladder_last_n_def ladder_n_of_cut ladder_n_prev_bound le_neq_implies_less length_E_eq_n' length_L' length_greater_0_conv less_trans u_less_v v_def) qed qed } note ladder_n_pairwise_bound = this have is_ladder_E_L': "is_ladder E L'" apply (auto simp add: is_ladder_def ladder_last_n_L') using length_L_nonzero length_L' apply simp using ladder_n_bound apply blast using ladder_n_pairwise_bound by blast { fix index :: nat assume index_bound: "index + 1 < length L" then have index_le: "index < length L" by arith from index_bound have len_L_minus_1: "length L - 1 \ 0" by arith obtain m where m: "m = ladder_n L index" by blast from LeftDerivationLadder_propagate[OF LDLadder ladder_i_in_\ m index_le] obtain \ where \: "LeftDerivation \ (take m D) \ \ ladder_\ (\ @ \) D L index = \ @ \ \ ladder_j L index < length \" using index_bound by auto have L'_Derive: "ladder_\ \ E L' index = Derive \ (take (ladder_n L' index) E)" by (simp add: ladder_\_def) have ladder_n_L'_eq_L: "ladder_n L' index = ladder_n L index" using L' index_bound ladder_n_of_cut length_L_nonzero by auto have "ladder_prev_n L (length L - 1) < n'" using finish len_L_minus_1 by blast then have n'_is_upper_bound: "ladder_n L (length L - 2) < n'" using index_bound by (metis diff_diff_left ladder_prev_n_def len_L_minus_1 one_add_one) have index_upper_bound: "index \ length L - 2" using index_bound by linarith have ladder_n_upper_bound: "ladder_n L index \ ladder_n L (length L - 2)" apply (rule_tac ladder_n_monotone) apply (rule_tac is_ladder_D_L) apply (rule index_upper_bound) using length_L_nonzero by linarith with n'_is_upper_bound have m_le_n': "m \ n'" using dual_order.strict_implies_order le_less_trans m by linarith then have "take m E = take m D" by (metis E le_take_same length_E_eq_n' order_refl take_all) then have take_helper: "(take (ladder_n L' index) E) = take m D" by (simp add: ladder_n_L'_eq_L m) then have Derive_eq_\: "Derive \ (take (ladder_n L' index) E) = \" by (simp add: Derive LeftDerivation_implies_Derivation \) then have t1: "ladder_\ (\@\) D L index = (ladder_\ \ E L' index) @ \" by (simp add: L'_Derive \) have \_eq: "\ = ladder_\ \ E L' index" by (simp add: Derive_eq_\ L'_Derive) then have t2: "LeftDerivation \ (take (ladder_n L index) D) (ladder_\ \ E L' index)" using \ m by blast have t3: "(take (ladder_n L' index) E) = take (ladder_n L index) D" by (simp add: m take_helper) have t4: "ladder_j L index < length (ladder_\ \ E L' index)" using \ \_eq by blast have t5: "E ! (ladder_n L' index) = D ! (ladder_n L index)" using E ladder_n_L'_eq_L ladder_n_upper_bound n'_is_upper_bound by auto note t = t1 t2 t3 t4 t5 } note ladder_early_stage = this { fix index :: nat assume index_bound: "index < length L" have i: "ladder_i L' index = ladder_i L index" using L' ladder_i_cut by (simp add: index_bound) have j: "ladder_j L' index = ladder_j L index" using L' ladder_j_cut by (simp add: index_bound) have ix: "index > 0 \ ladder_ix L' index = ladder_ix L index" using L' ladder_ix_cut by (simp add: index_bound) have \: "ladder_\ (\@\) D L index = (ladder_\ \ E L' index) @ \" by (simp add: index_bound ladder_\_def ladder_early_stage(1)) have i_bound: "index > 0 \ ladder_i L' index < length (ladder_\ \ E L' index)" using i index_bound ladder_\_def ladder_early_stage(4) ladder_i_def by auto note ij = i j ix \ i_bound } note ladder_every_stage = this { fix u :: nat fix v :: nat assume u_le_v: "u \ v" assume v_bound: "v < length L" have "ladder_n L u \ ladder_n L v" using is_ladder_D_L ladder_n_monotone u_le_v v_bound by blast } note ladder_L_n_pairwise_le = this { fix index :: nat assume index_lower_bound: "index > 0" assume index_upper_bound: "index + 1 < length L" note derivation = ladder_early_stage(2) have derivation1: "LeftDerivation \ (take (ladder_n L (index - Suc 0)) D) (ladder_\ \ E L' (index - Suc 0))" using derivation[of "index - Suc 0"] index_lower_bound index_upper_bound using One_nat_def Suc_diff_1 Suc_eq_plus1 le_less_trans lessI less_or_eq_imp_le by linarith have derivation2: "LeftDerivation \ (take (ladder_n L index) D) (ladder_\ \ E L' index)" using derivation[of index] index_upper_bound by blast have ladder_\_is_\[symmetric]: "ladder_\ \ E L' (index - Suc 0) = ladder_\ \ E L' index" using index_lower_bound ladder_\_def by auto have ladder_n_le: "ladder_n L (index - Suc 0) \ ladder_n L index" apply (rule_tac ladder_L_n_pairwise_le) apply arith using index_upper_bound by arith from LeftDerivation_from_in_between[OF derivation1 derivation2 ladder_n_le] ladder_\_is_\ have "LeftDerivation (ladder_\ \ E L' index) (drop (ladder_n L' (index - Suc 0)) (take (ladder_n L' index) E)) (ladder_\ \ E L' index)" by (metis L' Suc_leI add_lessD1 index_lower_bound index_upper_bound ladder_early_stage(3) ladder_n_of_cut le_add_diff_inverse2 length_L_nonzero length_greater_0_conv less_diff_conv) } note LeftDerivation_delta_early = this have LeftDerivationFix_\_0: "LeftDerivationFix \ (ladder_i L' 0) (take (ladder_n L' 0) E) (ladder_j L' 0) (ladder_\ \ E L' 0)" proof - have ldfix: "LeftDerivationFix (\@\) (ladder_i L 0) (take (ladder_n L 0) D) (ladder_j L 0) (ladder_\ (\@\) D L 0)" using LDLadder LeftDerivationLadder_def by blast have ladder_i_L': "ladder_i L' 0 = ladder_i L 0" using L' ladder_i_of_cut_at_0 length_L_nonzero by blast have ladder_j_L': "ladder_j L' 0 = ladder_j L 0" by (metis L' ladder_cut_def ladder_j_def ladder_last_j_def ladder_last_j_of_cut length_L' length_greater_0_conv nth_list_update_neq) have "length L = 1 \ length L > 1" using length_L_nonzero by linarith then show ?thesis proof (induct rule: disjCases2) case 1 from 1 have ladder_n_L'_0: "ladder_n L' 0 = n'" using diff_is_0_eq' ladder_last_n_L' ladder_last_n_def length_E_eq_n' length_L' less_or_eq_imp_le by auto have take_n'_E: "take n' E = E" by (simp add: length_E_eq_n') from ladder_n_L'_0 take_n'_E have take_ladder_n_L': "take (ladder_n L' 0) E = E" by auto have "ladder_n L 0 = length D" by (simp add: "1.hyps" length_D_eq_n n_eq_ladder_n) then have take_ladder_n_L_0: "take (ladder_n L 0) D = D" by simp have ladder_\_\: "ladder_\ \ E L' 0 = \" apply (simp add: ladder_\_def take_ladder_n_L') by (simp add: Derive E LeftDerivation_implies_Derivation finish) have ladder_j_in_\: "ladder_j L 0 < length \" using finish "1.hyps" by auto have ldfix_1: "LeftDerivationFix (\@\) (ladder_i L 0) D (ladder_j L 0) (\@\')" using "1.hyps" \ \_ladder ldfix take_ladder_n_L_0 by auto then have "LeftDerivationFix \ (ladder_i L 0) E (ladder_j L 0) \" by (simp add: E LeftDerivationFix_cut_appendix finish ladder_i_in_\ ladder_j_in_\ length_D_eq_n) then show ?case by (simp add: ladder_i_L' ladder_j_L' take_ladder_n_L' ladder_\_\) next case 2 have h: "0 + 1 < length L" using "2.hyps" by auto note stage = ladder_early_stage[of 0, OF h] have ldfix0: "LeftDerivationFix (\@\) (ladder_i L 0) (take (ladder_n L 0) D) (ladder_j L 0) ((ladder_\ \ E L' 0) @ \)" using ladder_i_L' ladder_j_L' ldfix stage(1) stage(3) by auto from LeftDerivationFix_cut_appendix'[OF ldfix0 stage(2) ladder_i_in_\ stage(4)] show ?case by (simp add: ladder_i_L' ladder_j_L' stage(3)) qed qed { fix index :: nat assume index_bounds: "1 \ index \ index + 1 < length L" have introsAt_appendix: "LeftDerivationIntrosAt (\@\) D L index" using LDLadder LeftDerivationIntros_def LeftDerivationLadder_def add_lessD1 index_bounds by blast have index_plus_1_upper_bound: "index + 1 < length L" using index_bounds by arith note early_stage = ladder_early_stage[of index, OF index_plus_1_upper_bound] have ladder_i_L_index_eq_fst: "ladder_i L index = fst (D ! ladder_n L (index - Suc 0))" using introsAt_appendix LeftDerivationIntrosAt_def index_bounds by (metis One_nat_def) have E_at_D_at: "(E ! ladder_n L' (index - Suc 0)) = (D ! ladder_n L (index - Suc 0))" using ladder_early_stage[of "index - Suc 0"] by (metis One_nat_def add_lessD1 index_bounds le_add_diff_inverse2) then have ladder_i_L'_index_eq_fst: "ladder_i L' index = fst (E ! ladder_n L' (index - Suc 0))" using index_bounds ladder_i_L_index_eq_fst ladder_every_stage(1) le_add1 le_less_trans by auto have same_derivation: "(drop (Suc (ladder_n L' (index - Suc 0))) (take (ladder_n L' index) E)) = (drop (Suc (ladder_n L (index - Suc 0))) (take (ladder_n L index) D))" using L' early_stage(3) index_bounds ladder_n_of_cut length_L_nonzero by auto have deriv_split: "(drop (ladder_n L' (index - Suc 0)) (take (ladder_n L' index) E)) = ((ladder_i L' index, snd (E ! ladder_n L' (index - Suc 0))) # drop (Suc (ladder_n L' (index - Suc 0))) (take (ladder_n L' index) E))" by (metis Cons_nth_drop_Suc One_nat_def Suc_le_lessD add_lessD1 diff_Suc_less index_bounds ladder_i_L'_index_eq_fst ladder_n_bound ladder_n_pairwise_bound length_L' length_take min.absorb2 nth_take prod.collapse) have "LeftDerivationIntrosAt \ E L' index" apply (auto simp add: LeftDerivationIntrosAt_def Let_def) using ladder_i_L'_index_eq_fst apply blast apply (rule_tac LeftDerivationIntro_cut_appendix'[where \=\ and \' = \]) apply (metis E_at_D_at LeftDerivationIntrosAt_def One_nat_def Suc_le_lessD add_lessD1 early_stage(1) index_bounds introsAt_appendix ladder_every_stage(2) ladder_every_stage(3) ladder_every_stage(4) ladder_i_L'_index_eq_fst same_derivation) defer 1 using index_bounds ladder_every_stage(5) apply auto[1] using early_stage(4) index_bounds ladder_every_stage(2) apply auto[1] using LeftDerivation_delta_early deriv_split by (metis One_nat_def Suc_le_eq index_bounds) } note LeftDerivationIntrosAt_early = this { fix index :: nat assume index_bounds: "1 \ index \ index + 1 = length L" have introsAt_appendix: "LeftDerivationIntrosAt (\@\) D L index" using LDLadder LeftDerivationIntros_def LeftDerivationLadder_def add_lessD1 index_bounds by (metis Suc_eq_plus1 not_less_eq) have ladder_i_L_index_eq_fst: "ladder_i L index = fst (D ! ladder_n L (index - Suc 0))" using introsAt_appendix LeftDerivationIntrosAt_def index_bounds by (metis One_nat_def) have E_at_D_at: "(E ! ladder_n L' (index - Suc 0)) = (D ! ladder_n L (index - Suc 0))" using ladder_early_stage[of "index - Suc 0"] by (metis One_nat_def Suc_eq_plus1 index_bounds le_add_diff_inverse2 lessI) then have ladder_i_L'_index_eq_fst: "ladder_i L' index = fst (E ! ladder_n L' (index - Suc 0))" using index_bounds ladder_i_L_index_eq_fst ladder_every_stage(1) le_add1 le_less_trans by auto obtain D' where D': "D' = (drop (Suc (ladder_n L (index - Suc 0))) (take (ladder_n L index) D))" by blast obtain k where k: "k = (ladder_n L' index) - (Suc (ladder_n L' (index - Suc 0)))" by blast have ladder_n_L'_index: "ladder_n L' index = length E" by (metis diff_add_inverse2 index_bounds ladder_last_n_L' ladder_last_n_def length_L') have take_is_E: "take (ladder_n L' index) E = E" by (simp add: ladder_n_L'_index) have ladder_n_L_index: "ladder_n L index = length D" by (metis diff_add_inverse2 index_bounds length_D_eq_n n_eq_ladder_n) have take_is_D: "take (ladder_n L index) D = D" by (simp add: ladder_n_L_index) have write_as_take_k_D': "(drop (Suc (ladder_n L' (index - Suc 0))) E) = take k D'" using take_is_D by (metis D' E L' One_nat_def Suc_le_lessD add_diff_cancel_right' diff_Suc_less drop_take index_bounds k ladder_n_L'_index ladder_n_of_cut length_E_eq_n' length_L_nonzero length_greater_0_conv) have k_bound: "k \ length D'" by (metis le_iff_add append_take_drop_id k ladder_n_L'_index length_append length_drop write_as_take_k_D') have D'_alt: "D' = drop (Suc (ladder_n L (index - Suc 0))) D" by (simp add: D' take_is_D) have "LeftDerivationIntrosAt \ E L' index" apply (auto simp add: LeftDerivationIntrosAt_def Let_def) using ladder_i_L'_index_eq_fst apply blast apply (subst take_is_E) apply (subst write_as_take_k_D') apply (rule_tac LeftDerivationIntro_cut_appendix[where \=\ and \' = \']) apply (metis D' Derive E E_at_D_at LeftDerivationIntrosAt_def LeftDerivation_implies_Derivation One_nat_def Suc_le_lessD add_diff_cancel_right' diff_Suc_less finish index_bounds introsAt_appendix ladder_\_def ladder_every_stage(2) ladder_every_stage(3) ladder_every_stage(4) ladder_i_L'_index_eq_fst length_L_nonzero take_is_E) apply (metis Cons_nth_drop_Suc E L' LeftDerivation_from_in_between LeftDerivation_take_derive One_nat_def Suc_le_lessD add_diff_cancel_right' diff_Suc_less finish index_bounds ladder_\_def ladder_\_def ladder_i_L'_index_eq_fst ladder_n_L'_index ladder_n_of_cut ladder_prev_n_def length_E_eq_n' length_L_nonzero less_imp_le_nat less_numeral_extra(3) list.size(3) prod.collapse take_is_E write_as_take_k_D') using k_bound apply blast using D'_alt apply (metis (no_types, lifting) Derive E L' LeftDerivation_implies_Derivation One_nat_def Suc_leI Suc_le_lessD add_diff_cancel_right' diff_Suc_less drop_drop finish index_bounds k ladder_\_def ladder_n_L'_index ladder_n_of_cut ladder_prev_n_def le_add_diff_inverse2 length_E_eq_n' length_L_nonzero length_greater_0_conv less_not_refl2 take_is_E) using index_bounds ladder_every_stage(5) apply auto[1] by (metis Derive E LeftDerivation_implies_Derivation One_nat_def add_diff_cancel_right' diff_Suc_less finish index_bounds ladder_\_def ladder_every_stage(2) length_L_nonzero take_is_E) } note LeftDerivationIntrosAt_last = this have ladder_E_L': "LeftDerivationLadder \ E L' \" apply (auto simp add: LeftDerivationLadder_def) using finish E apply blast using is_ladder_E_L' apply blast using LeftDerivationFix_\_0 apply blast using LeftDerivationIntros_def LeftDerivationIntrosAt_early LeftDerivationIntrosAt_last by (metis Suc_eq_plus1 Suc_leI le_neq_implies_less length_L') show ?thesis apply (rule_tac x=E in exI) apply (rule_tac x=F in exI) apply (rule_tac x=\ in exI) apply (rule_tac x=\' in exI) apply (rule_tac x=L' in exI) apply auto using E F apply simp apply (rule \) using ladder_E_L' apply blast using F finish apply blast using F finish apply blast by (rule L') qed theorem LeftDerivationLadder_cut_appendix: assumes LDLadder: "LeftDerivationLadder (\@\) D L \" assumes ladder_i_in_\: "ladder_i L 0 < length \" shows "\ E F \1 \2 L'. D = E@F \ \ = \1 @ \2 \ LeftDerivationLadder \ E L' \1 \ derivation_ge F (length \1) \ LeftDerivation \ (derivation_shift F (length \1) 0) \2 \ length L' = length L \ ladder_i L' 0 = ladder_i L 0 \ ladder_last_j L' = ladder_last_j L" proof - from LeftDerivationLadder_cut_appendix_helper[OF LDLadder ladder_i_in_\] obtain E F \1 \2 L' where helper: "D = E @ F \ \ = \1 @ \2 \ LeftDerivationLadder \ E L' \1 \ derivation_ge F (length \1) \ LeftDerivation \ (derivation_shift F (length \1) 0) \2 \ L' = ladder_cut L (length E)" by blast show ?thesis apply (rule_tac x=E in exI) apply (rule_tac x=F in exI) apply (rule_tac x=\1 in exI) apply (rule_tac x=\2 in exI) apply (rule_tac x=L' in exI) using helper LDLadder LeftDerivationLadder_def is_ladder_def ladder_i_of_cut_at_0 ladder_last_j_of_cut length_ladder_cut by force qed definition ladder_stepdown_diff :: "ladder \ nat" where "ladder_stepdown_diff L = Suc (ladder_n L 0)" definition ladder_stepdown_\_0 :: "sentence \ derivation \ ladder \ sentence" where "ladder_stepdown_\_0 a D L = Derive a (take (ladder_stepdown_diff L) D)" lemma LeftDerivationIntro_LeftDerives1: assumes "LeftDerivationIntro \ i r ix D j \" assumes "splits_at \ i a1 A a2" shows "LeftDerives1 \ i r (a1@(snd r)@a2)" by (metis LeftDerivationIntro_def LeftDerivationIntro_examine_rule LeftDerivationIntro_is_sentence LeftDerives1_def assms(1) assms(2) prod.collapse splits_at_implies_Derives1) lemma LeftDerives1_Derive: assumes "LeftDerives1 \ i r \" shows "Derive \ [(i, r)] = \" by (metis Derive LeftDerivation.simps(1) LeftDerivation_LeftDerives1 LeftDerivation_implies_Derivation append_Nil assms) lemma ladder_stepdown_\_0_altdef: assumes ladder: "LeftDerivationLadder \ D L \" assumes length_L: "length L > 1" assumes split: "splits_at (ladder_\ \ D L 1) (ladder_i L 1) a1 A a2" shows "ladder_stepdown_\_0 \ D L = a1 @ (snd (snd (D ! (ladder_n L 0)))) @ a2" proof - have 1: "ladder_\ \ D L 1 = Derive \ (take (ladder_n L 0) D)" by (simp add: ladder_\_def ladder_\_def) obtain rule where rule: "rule = snd (D ! (ladder_n L 0))" by blast have "\ E \. LeftDerivationIntro (ladder_\ \ D L 1) (ladder_i L 1) rule (ladder_ix L 1) E (ladder_j L 1) \" by (metis LeftDerivationIntrosAt_def LeftDerivationIntros_def LeftDerivationLadder_def One_nat_def diff_Suc_1 ladder length_L order_refl rule) then obtain E \ where intro: "LeftDerivationIntro (ladder_\ \ D L 1) (ladder_i L 1) rule (ladder_ix L 1) E (ladder_j L 1) \" by blast then have 2: "LeftDerives1 (ladder_\ \ D L 1) (ladder_i L 1) rule (a1@(snd rule)@a2)" using LeftDerivationIntro_LeftDerives1 split by blast have fst_D: "fst (D ! (ladder_n L 0)) = ladder_i L 1" by (metis LeftDerivationIntrosAt_def LeftDerivationIntros_def LeftDerivationLadder_def One_nat_def diff_Suc_1 ladder le_numeral_extra(4) length_L) have derive_derive: "Derive \ (take (Suc (ladder_n L 0)) D) = Derive (Derive \ (take (ladder_n L 0) D)) [D ! (ladder_n L 0)]" proof - have f1: "Derivation \ (take (Suc (ladder_n L 0)) D) (Derive \ (take (Suc (ladder_n L 0)) D))" using Derivation_take_derive LeftDerivationLadder_def LeftDerivation_implies_Derivation ladder by blast have f2: "length L - 1 < length L" using length_L by linarith have "0 < length L - 1" using length_L by linarith then have f3: "take (Suc (ladder_n L 0)) D = take (ladder_n L 0) D @ [D ! ladder_n L 0]" using f2 by (metis (full_types) LeftDerivationLadder_def is_ladder_def ladder ladder_last_n_def take_Suc_conv_app_nth) obtain sss :: "symbol list \ (nat \ symbol \ symbol list) list \ (nat \ symbol \ symbol list) list \ symbol list \ symbol list" where "\x0 x1 x2 x3. (\v4. Derivation x3 x2 v4 \ Derivation v4 x1 x0) = (Derivation x3 x2 (sss x0 x1 x2 x3) \ Derivation (sss x0 x1 x2 x3) x1 x0)" by moura then show ?thesis using f3 f1 Derivation_append Derive by auto qed then have 3: "ladder_stepdown_\_0 \ D L = Derive (ladder_\ \ D L 1) [D ! (ladder_n L 0)]" using 1 by (simp add: ladder_stepdown_\_0_def ladder_stepdown_diff_def) have 4: "D ! (ladder_n L 0) = (ladder_i L 1, rule)" using rule fst_D by (metis prod.collapse) then show ?thesis using 2 3 4 LeftDerives1_Derive snd_conv by auto qed lemma ladder_i_0_bound: assumes ld: "LeftDerivationLadder \ D L \" shows "ladder_i L 0 < length \" proof - have "LeftDerivationFix \ (ladder_i L 0) (take (ladder_n L 0) D) (ladder_j L 0) (ladder_\ \ D L 0)" using ld LeftDerivationLadder_def by simp then show ?thesis using LeftDerivationFix_def by simp qed lemma ladder_j_bound: assumes ld: "LeftDerivationLadder \ D L \" assumes index_bound: "index < length L" shows "ladder_j L index < length (ladder_\ \ D L index)" proof - have ld': "LeftDerivationLadder (\@[]) D L \" using ld by simp have ladder_i_0: "ladder_i L 0 < length \" using ladder_i_0_bound ld by auto obtain n where n: "n = ladder_n L index" by blast note propagate = LeftDerivationLadder_propagate[OF ld' ladder_i_0 n index_bound] from index_bound have "index + 1 < length L \ index + 1 = length L" by arith then show ?thesis proof (induct rule: disjCases2) case 1 then have "\\. LeftDerivation \ (take n D) \ \ ladder_\ (\ @ []) D L index = \ @ [] \ ladder_j L index < length \" using propagate by auto then show ?case by auto next case 2 then have " \n' \ \'. (index = 0 \ ladder_prev_n L index < n') \ n' \ n \ LeftDerivation \ (take n' D) \ \ LeftDerivation (\ @ []) (take n' D) (\ @ []) \ derivation_ge (drop n' D) (length \) \ LeftDerivation [] (derivation_shift (drop n' D) (length \) 0) \' \ ladder_\ (\ @ []) D L index = \ @ \' \ ladder_j L index < length \" using propagate by auto then show ?case by auto qed qed lemma ladder_last_j_bound: assumes ld: "LeftDerivationLadder \ D L \" shows "ladder_last_j L < length \" proof - have "length L - 1 < length L" using LeftDerivationLadder_def assms is_ladder_def by auto from ladder_j_bound[OF ld this] show ?thesis by (metis Derive LeftDerivationLadder_def LeftDerivation_implies_Derivation One_nat_def is_ladder_def ladder_last_j_def last_ladder_\ ld) qed fun ladder_shift_n :: "nat \ ladder \ ladder" where "ladder_shift_n N [] = []" | "ladder_shift_n N ((n, j, i)#L) = ((n - N, j, i)#(ladder_shift_n N L))" fun ladder_stepdown :: "ladder \ ladder" where "ladder_stepdown [] = undefined" | "ladder_stepdown [v] = undefined" | "ladder_stepdown ((n0, j0, i0)#(n1, j1, ix1)#L) = (n1 - Suc n0, j1, j0 + ix1) # (ladder_shift_n (Suc n0) L)" lemma ladder_shift_n_length: "length (ladder_shift_n N L) = length L" by (induct L, auto) lemma ladder_stepdown_prepare: assumes "length L > 1" shows "L = (ladder_n L 0, ladder_j L 0, ladder_i L 0)# (ladder_n L 1, ladder_j L 1, ladder_ix L 1)#(drop 2 L)" proof - have "\ n0 j0 i0 n1 j1 ix1 L'. L = ((n0, j0, i0)#(n1, j1, ix1)#L')" by (metis One_nat_def Suc_eq_plus1 assms ladder_stepdown.cases less_not_refl list.size(3) list.size(4) not_less0) then obtain n0 j0 i0 n1 j1 ix1 L' where L': "L = ((n0, j0, i0)#(n1, j1, ix1)#L')" by blast have n0: "n0 = ladder_n L 0" using L' by (auto simp add: ladder_n_def deriv_n_def) show ?thesis using L' by (auto simp add: ladder_n_def deriv_n_def ladder_j_def deriv_j_def ladder_i_def deriv_i_def ladder_ix_def deriv_ix_def) qed lemma ladder_stepdown_length: assumes "length L > 1" shows "length (ladder_stepdown L) = length L - 1" apply (subst ladder_stepdown_prepare[OF assms(1)]) apply (simp add: ladder_shift_n_length) using assms apply arith done lemma ladder_stepdown_i_0: assumes "length L > 1" shows "ladder_i (ladder_stepdown L) 0 = ladder_i L 1 + ladder_ix L 1" using ladder_stepdown_prepare[OF assms] ladder_i_def ladder_j_def deriv_j_def by (metis One_nat_def deriv_i_def diff_Suc_1 ladder_stepdown.simps(3) list.sel(1) snd_conv zero_neq_one) lemma ladder_shift_n_cons: "ladder_shift_n N (x#L) = (fst x - N, snd x)#(ladder_shift_n N L)" apply (induct L) by (cases x, simp)+ lemma ladder_shift_n_drop: "ladder_shift_n N (drop n L) = drop n (ladder_shift_n N L)" proof (induct L arbitrary: n) case Nil then show ?case by simp next case (Cons x L) show ?case proof (cases n) case 0 then show ?thesis by simp next case (Suc n) then show ?thesis by (simp add: ladder_shift_n_cons Cons) qed qed lemma drop_2_shift: assumes "index > 0" assumes "length L > 1" shows "drop 2 L ! (index - Suc 0) = L ! Suc index" proof - define l1 l2 and L' where "l1 = L ! 0" "l2 = L ! 1" and "L' = drop 2 L" with \length L > 1\ have "L = l1 # l2 # L'" by (cases L) (auto simp add: neq_Nil_conv) with \index > 0\ show ?thesis by simp qed lemma ladder_shift_n_at: "index < length L \ (ladder_shift_n N L) ! index = (fst (L ! index) - N, snd (L ! index))" proof (induct L arbitrary: index) case Nil then show ?case by auto next case (Cons x L) show ?case apply (simp add: ladder_shift_n_cons) apply (cases index) apply (auto) apply (rule_tac Cons(1)) using Cons(2) by auto qed lemma ladder_stepdown_j: assumes length_L_greater_1: "length L > 1" assumes L': "L' = ladder_stepdown L" assumes index_bound: "index < length L'" shows "ladder_j L' index = ladder_j L (Suc index)" proof - note L_prepare = ladder_stepdown_prepare[OF length_L_greater_1] have ladder_stepdown_L_def: "ladder_stepdown L = ((ladder_n L (Suc 0) - Suc (ladder_n L 0), ladder_j L (Suc 0), ladder_j L 0 + ladder_ix L (Suc 0)) # ladder_shift_n (Suc (ladder_n L 0)) (drop 2 L))" by (subst L_prepare, simp) have "index = 0 \ index > 0" by arith then show "ladder_j L' index = ladder_j L (Suc index)" proof (induct rule: disjCases2) case 1 show ?case by (simp add: L' ladder_stepdown_L_def 1 ladder_j_def deriv_j_def) next case 2 have index_bound': "Suc index < length L" using index_bound L' ladder_stepdown_length length_L_greater_1 by auto show ?case apply (simp add: L' ladder_stepdown_L_def 2 ladder_j_def ladder_shift_n_drop drop_2_shift) apply (subst drop_2_shift) apply (simp add: 2) using length_L_greater_1 apply (simp add: ladder_shift_n_length) apply (simp add: deriv_j_def) apply (simp add: ladder_shift_n_at[OF index_bound']) done qed qed lemma ladder_stepdown_last_j: assumes length_L_greater_1: "length L > 1" shows "ladder_last_j (ladder_stepdown L) = ladder_last_j L" using ladder_stepdown_j Suc_diff_Suc diff_Suc_1 ladder_last_j_def ladder_stepdown_length length_L_greater_1 lessI by auto lemma ladder_stepdown_n: assumes length_L_greater_1: "length L > 1" assumes L': "L' = ladder_stepdown L" assumes index_bound: "index < length L'" shows "ladder_n L' index = ladder_n L (Suc index) - ladder_stepdown_diff L" proof - note L_prepare = ladder_stepdown_prepare[OF length_L_greater_1] have ladder_stepdown_L_def: "ladder_stepdown L = ((ladder_n L (Suc 0) - Suc (ladder_n L 0), ladder_j L (Suc 0), ladder_j L 0 + ladder_ix L (Suc 0)) # ladder_shift_n (Suc (ladder_n L 0)) (drop 2 L))" by (subst L_prepare, simp) have "index = 0 \ index > 0" by arith then show "ladder_n L' index = ladder_n L (Suc index) - ladder_stepdown_diff L" proof (induct rule: disjCases2) case 1 show ?case by (simp add: L' ladder_stepdown_L_def 1 ladder_n_def deriv_n_def ladder_stepdown_diff_def) next case 2 have index_bound': "Suc index < length L" using index_bound L' ladder_stepdown_length length_L_greater_1 by auto show ?case apply (simp add: L' ladder_stepdown_L_def 2 ladder_n_def ladder_shift_n_drop drop_2_shift ladder_stepdown_diff_def) apply (subst drop_2_shift) apply (simp add: 2) using length_L_greater_1 apply (simp add: ladder_shift_n_length) apply (simp add: deriv_n_def) apply (simp add: ladder_shift_n_at[OF index_bound']) done qed qed lemma ladder_stepdown_ix: assumes length_L_greater_1: "length L > 1" assumes L': "L' = ladder_stepdown L" assumes index_lower_bound: "0 < index" assumes index_upper_bound: "index < length L'" shows "ladder_ix L' index = ladder_ix L (Suc index)" proof - note L_prepare = ladder_stepdown_prepare[OF length_L_greater_1] have ladder_stepdown_L_def: "ladder_stepdown L = ((ladder_n L (Suc 0) - Suc (ladder_n L 0), ladder_j L (Suc 0), ladder_j L 0 + ladder_ix L (Suc 0)) # ladder_shift_n (Suc (ladder_n L 0)) (drop 2 L))" by (subst L_prepare, simp) have index_bound': "Suc index < length L" using index_upper_bound L' ladder_stepdown_length length_L_greater_1 by auto show ?thesis apply (simp add: L' ladder_stepdown_L_def index_lower_bound ladder_ix_def ladder_shift_n_drop) apply (subst drop_2_shift) apply (simp add: index_lower_bound) using length_L_greater_1 apply (simp add: ladder_shift_n_length) apply (simp add: deriv_ix_def) apply (simp add: ladder_shift_n_at[OF index_bound']) using index_lower_bound by arith qed lemma Derive_Derive: assumes "Derivation \ (D@E) \" shows "Derive (Derive \ D) E = Derive \ (D@E)" using Derivation_append Derive assms by fastforce lemma drop_at_shift: assumes "n \ index" assumes "index < length D" shows "drop n D ! (index - n) = D ! index" using assms(1) assms(2) by auto theorem LeftDerivationLadder_stepdown: assumes ldl: "LeftDerivationLadder \ D L \" assumes length_L: "length L > 1" shows "\ L'. LeftDerivationLadder (ladder_stepdown_\_0 \ D L) (drop (ladder_stepdown_diff L) D) L' \ \ length L' = length L - 1 \ ladder_i L' 0 = ladder_i L 1 + ladder_ix L 1 \ ladder_last_j L' = ladder_last_j L" proof - obtain L' where L': "L' = ladder_stepdown L" by blast have ldl1: "LeftDerivation (ladder_stepdown_\_0 \ D L) (drop (ladder_stepdown_diff L) D) \" proof - have D_split: "D = (take (ladder_stepdown_diff L) D) @ (drop (ladder_stepdown_diff L) D)" by simp show ?thesis using D_split ldl proof - obtain sss :: "symbol list \ (nat \ symbol \ symbol list) list \ (nat \ symbol \ symbol list) list \ symbol list \ symbol list" where "\x0 x1 x2 x3. (\v4. LeftDerivation x3 x2 v4 \ LeftDerivation v4 x1 x0) = (LeftDerivation x3 x2 (sss x0 x1 x2 x3) \ LeftDerivation (sss x0 x1 x2 x3) x1 x0)" by moura then have "(\ LeftDerivation \ (take (ladder_stepdown_diff L) D @ drop (ladder_stepdown_diff L) D) \ \ LeftDerivation \ (take (ladder_stepdown_diff L) D) (sss \ (drop (ladder_stepdown_diff L) D) (take (ladder_stepdown_diff L) D) \) \ LeftDerivation (sss \ (drop (ladder_stepdown_diff L) D) (take (ladder_stepdown_diff L) D) \) (drop (ladder_stepdown_diff L) D) \) \ (LeftDerivation \ (take (ladder_stepdown_diff L) D @ drop (ladder_stepdown_diff L) D) \ \ (\ss. \ LeftDerivation \ (take (ladder_stepdown_diff L) D) ss \ \ LeftDerivation ss (drop (ladder_stepdown_diff L) D) \))" using LeftDerivation_append by blast then show ?thesis by (metis (no_types) D_split Derivation_take_derive Derivation_unique_dest LeftDerivationLadder_def LeftDerivation_implies_Derivation ladder_stepdown_\_0_def ldl) qed qed have L'_nonempty: "L' \ []" using L' ladder_stepdown_length length_L by fastforce { fix u :: nat assume u': "u < length L'" then have Suc_u: "Suc u < length L" using L' ladder_stepdown_length length_L by auto have "ladder_n L (Suc u) \ length D" using ldl Suc_u by (simp add: LeftDerivationLadder_ladder_n_bound) then have "ladder_n L' u \ length D - ladder_stepdown_diff L" apply (subst ladder_stepdown_n[OF length_L L' u']) by auto } note is_ladder_prop1 = this { fix u :: nat fix v :: nat assume u_less_v: "u < v" assume v_L': "v < length L'" from u_less_v v_L' have u_L': "u < length L'" by arith have "ladder_n L (Suc u) < ladder_n L (Suc v)" using ldl by (metis (no_types, lifting) L' LeftDerivationLadder_def One_nat_def Suc_diff_1 Suc_lessD Suc_mono is_ladder_def ladder_stepdown_length length_L u_less_v v_L') then have "ladder_n L' u < ladder_n L' v" apply (simp add: ladder_stepdown_n[OF length_L L'] u_L' v_L') by (metis (no_types, lifting) L' LeftDerivationLadder_def Suc_eq_plus1 Suc_leI diff_less_mono is_ladder_def ladder_stepdown_diff_def ladder_stepdown_length ldl length_L less_diff_conv u_L' zero_less_Suc) } note is_ladder_prop2 = this have is_ladder_L': "is_ladder (drop (ladder_stepdown_diff L) D) L'" apply (auto simp add: is_ladder_def) using L'_nonempty apply blast using is_ladder_prop1 apply blast using is_ladder_prop2 apply blast using ladder_last_n_def ladder_stepdown_n L' LeftDerivationLadder_def Suc_diff_Suc diff_Suc_1 ladder_n_last_is_length ladder_stepdown_length ldl length_L lessI by auto have ldfix: "LeftDerivationFix (ladder_stepdown_\_0 \ D L) (ladder_i L' 0) (take (ladder_n L' 0) (drop (ladder_stepdown_diff L) D)) (ladder_j L' 0) (ladder_\ (ladder_stepdown_\_0 \ D L) (drop (ladder_stepdown_diff L) D) L' 0)" proof - have introsAt_L_1: "LeftDerivationIntrosAt \ D L 1" using LeftDerivationIntros_def LeftDerivationLadder_def ldl length_L by blast thm LeftDerivationIntrosAt_def obtain n where n: "n = ladder_n L 0" by blast obtain m where m: "m = ladder_n L 1" by blast have "LeftDerivationIntro (ladder_\ \ D L 1) (ladder_i L 1) (snd (D ! n)) (ladder_ix L 1) (drop (Suc n) (take m D)) (ladder_j L 1) (ladder_\ \ D L 1)" using n m introsAt_L_1 by (metis LeftDerivationIntrosAt_def One_nat_def diff_Suc_1) from iffD1[OF LeftDerivationIntro_def this] obtain \ where \: "LeftDerives1 (ladder_\ \ D L 1) (ladder_i L 1) (snd (D ! n)) \ \ ladder_ix L 1 < length (snd (snd (D ! n))) \ snd (snd (D ! n)) ! ladder_ix L 1 = ladder_\ \ D L 1 ! ladder_j L 1 \ LeftDerivationFix \ (ladder_i L 1 + ladder_ix L 1) (drop (Suc n) (take m D)) (ladder_j L 1) (ladder_\ \ D L 1)" by blast have "\ = Derive (ladder_\ \ D L 1) [D ! n]" by (metis (no_types, opaque_lifting) LeftDerivationIntrosAt_def LeftDerives1_Derive \ cancel_comm_monoid_add_class.diff_cancel introsAt_L_1 n prod.collapse) then have \_def: "\ = ladder_stepdown_\_0 \ D L" proof - obtain sss :: "nat \ symbol list \ symbol list" and ss :: "nat \ symbol list \ symbol" and sssa :: "nat \ symbol list \ symbol list" where "\x2 x3. (\v4 v5 v6. splits_at x3 x2 v4 v5 v6) = splits_at x3 x2 (sss x2 x3) (ss x2 x3) (sssa x2 x3)" by moura then have f1: "\ssa n p ssb. \ Derives1 ssa n p ssb \ splits_at ssa n (sss n ssa) (ss n ssa) (sssa n ssa)" using splits_at_ex by presburger then have "\ = sss (ladder_i L 1) (ladder_\ \ D L 1) @ snd (snd (D ! n)) @ sssa (ladder_i L 1) (ladder_\ \ D L 1)" by (meson LeftDerives1_implies_Derives1 \ splits_at_combine_dest) then show ?thesis using f1 by (metis (no_types) LeftDerives1_implies_Derives1 \ ladder_stepdown_\_0_altdef ldl length_L n) qed have ladder_i_L'_0: "ladder_i L' 0 = ladder_i L 1 + ladder_ix L 1" using L' ladder_stepdown_i_0 length_L by blast have derivation_eq: "(take (ladder_n L' 0) (drop (ladder_stepdown_diff L) D)) = (drop (Suc n) (take m D))" using n m by (metis L' L'_nonempty One_nat_def drop_take ladder_stepdown_diff_def ladder_stepdown_n length_L length_greater_0_conv) have ladder_j_L'_0: "ladder_j L' 0 = ladder_j L 1" using L' L'_nonempty ladder_stepdown_j length_L by auto have ladder_\: "(ladder_\ (ladder_stepdown_\_0 \ D L) (drop (ladder_stepdown_diff L) D) L' 0) = ladder_\ \ D L 1" by (metis Derivation_take_derive Derivation_unique_dest LeftDerivationFix_def LeftDerivation_implies_Derivation \ \_def derivation_eq ladder_\_def ldl1) from \_def \ ladder_i_L'_0 derivation_eq ladder_j_L'_0 ladder_\ show ?thesis by auto qed { fix index :: nat assume index_lower_bound: "Suc 0 \ index" assume index_upper_bound: "index < length L'" then have Suc_index_upper_bound: "Suc index < length L" using L' Suc_diff_Suc Suc_less_eq diff_Suc_1 ladder_stepdown_length length_L less_Suc_eq by auto then have intros_at_Suc_index: "LeftDerivationIntrosAt \ D L (Suc index)" by (metis LeftDerivationIntros_def LeftDerivationLadder_def Suc_eq_plus1_left ldl le_add1) from iffD1[OF LeftDerivationIntrosAt_def this] have ldintro: "let \' = ladder_\ \ D L (Suc index); i = ladder_i L (Suc index); j = ladder_j L (Suc index); ix = ladder_ix L (Suc index); \ = ladder_\ \ D L (Suc index); n = ladder_n L (Suc index - 1); m = ladder_n L (Suc index); e = D ! n; E = drop (Suc n) (take m D) in i = fst e \ LeftDerivationIntro \' i (snd e) ix E j \" by blast have index_minus_Suc_0_bound: "index - Suc 0 < length L'" by (simp add: index_upper_bound less_imp_diff_less) note helpers = length_L L' index_minus_Suc_0_bound have ladder_i_L'_index: "ladder_i L' index = fst (drop (ladder_stepdown_diff L) D ! ladder_n L' (index - Suc 0))" apply (auto simp add: ladder_i_def) using index_lower_bound apply arith apply (simp add: ladder_stepdown_n[OF helpers] ladder_stepdown_j[OF helpers]) apply (subst drop_at_shift) using LeftDerivationLadder_def Suc_index_upper_bound Suc_leI Suc_lessD is_ladder_def ladder_stepdown_diff_def ldl apply presburger apply (metis LeftDerivationLadder_def One_nat_def Suc_eq_plus1 Suc_index_upper_bound add.commute add_diff_cancel_right' ladder_n_minus_1_bound ldl le_add1) by (metis LeftDerivationIntrosAt_def intros_at_Suc_index diff_Suc_1 ladder_i_def nat.simps(3)) have intro_at_index: "LeftDerivationIntro (ladder_\ (ladder_stepdown_\_0 \ D L) (drop (ladder_stepdown_diff L) D) L' index) (ladder_i L' index) (snd (drop (ladder_stepdown_diff L) D ! ladder_n L' (index - Suc 0))) (ladder_ix L' index) (drop (Suc (ladder_n L' (index - Suc 0))) (take (ladder_n L' index) (drop (ladder_stepdown_diff L) D))) (ladder_j L' index) (ladder_\ (ladder_stepdown_\_0 \ D L) (drop (ladder_stepdown_diff L) D) L' index)" proof - have arg1: "(ladder_\ (ladder_stepdown_\_0 \ D L) (drop (ladder_stepdown_diff L) D) L' index) = ladder_\ \ D L (Suc index)" apply (auto simp add: ladder_\_def ladder_\_def) using index_lower_bound apply arith apply (simp add: ladder_stepdown_n[OF helpers] ladder_stepdown_\_0_def) apply (subst Derive_Derive[where \="ladder_\ \ D L index"]) apply (metis (no_types, lifting) Derivation_take_derive LeftDerivationLadder_def LeftDerivation_implies_Derivation Suc_index_upper_bound Suc_leI Suc_lessD add.commute is_ladder_def ladder_\_def ladder_stepdown_diff_def ldl le_add_diff_inverse2 take_add) by (metis LeftDerivationLadder_def Suc_index_upper_bound Suc_leI Suc_lessD add.commute is_ladder_def ladder_stepdown_diff_def ldl le_add_diff_inverse2 take_add) have arg2: "ladder_i L' index = ladder_i L (Suc index)" using L' index_lower_bound index_minus_Suc_0_bound ladder_i_def ladder_stepdown_j length_L by auto obtain n where n: "n = ladder_n L (Suc index - 1)" by blast have arg3: "(snd (drop (ladder_stepdown_diff L) D ! ladder_n L' (index - Suc 0))) = snd (D ! n)" apply (simp add: ladder_stepdown_n[OF helpers] index_lower_bound) apply (subst drop_at_shift) using index_lower_bound apply (metis (no_types, opaque_lifting) L' LeftDerivationLadder_def One_nat_def Suc_eq_plus1 add.commute diff_Suc_1 index_upper_bound is_ladder_def ladder_stepdown_diff_def ladder_stepdown_length ldl le_add_diff_inverse2 length_L less_or_eq_imp_le n nat.simps(3) neq0_conv not_less not_less_eq_eq) using index_lower_bound apply (metis LeftDerivationLadder_def One_nat_def Suc_index_upper_bound Suc_le_lessD Suc_pred diff_Suc_1 ladder_n_minus_1_bound ldl le_imp_less_Suc less_imp_le) using index_lower_bound n by (simp add: Suc_diff_le) have arg4: "ladder_ix L' index = ladder_ix L (Suc index)" using ladder_stepdown_ix L' Suc_le_lessD index_lower_bound index_upper_bound length_L by auto obtain m where m: "m = ladder_n L (Suc index)" by blast have Suc_index_Suc: "Suc (index - Suc 0) = index" using index_lower_bound by arith have arg5: "(drop (Suc (ladder_n L' (index - Suc 0))) (take (ladder_n L' index) (drop (ladder_stepdown_diff L) D))) = drop (Suc n) (take m D)" apply (simp add: ladder_stepdown_n[OF helpers] ladder_stepdown_n[OF length_L L' index_upper_bound] n m Suc_index_Suc) by (metis (no_types, lifting) LeftDerivationLadder_def Suc_eq_plus1_left Suc_index_upper_bound Suc_leI Suc_le_lessD Suc_lessD drop_drop drop_take index_lower_bound is_ladder_def ladder_stepdown_diff_def ldl le_add_diff_inverse2) have arg6: "ladder_j L' index = ladder_j L (Suc index)" using L' index_upper_bound ladder_stepdown_j length_L by blast have arg7: "(ladder_\ (ladder_stepdown_\_0 \ D L) (drop (ladder_stepdown_diff L) D) L' index) = ladder_\ \ D L (Suc index)" apply (simp add: ladder_\_def) apply (simp add: ladder_stepdown_n[OF length_L L' index_upper_bound] ladder_stepdown_\_0_def) apply (subst Derive_Derive[where \="ladder_\ \ D L (Suc index)"]) apply (metis (no_types, lifting) L' LeftDerivationLadder_def LeftDerivation_implies_Derivation LeftDerivation_take_derive Suc_le_lessD add_diff_inverse_nat diff_is_0_eq index_lower_bound index_upper_bound is_ladder_L' is_ladder_def ladder_\_def ladder_stepdown_n ldl le_0_eq length_L less_numeral_extra(3) less_or_eq_imp_le take_add) by (metis (no_types, lifting) L' One_nat_def add_diff_inverse_nat diff_is_0_eq index_lower_bound index_upper_bound is_ladder_L' is_ladder_def ladder_stepdown_n le_0_eq le_neq_implies_less length_L less_numeral_extra(3) less_or_eq_imp_le take_add zero_less_one) from ldintro arg1 arg2 arg3 arg4 arg5 arg6 arg7 show ?thesis by (metis m n) qed have "LeftDerivationIntrosAt (ladder_stepdown_\_0 \ D L) (drop (ladder_stepdown_diff L) D) L' index" apply (auto simp add: LeftDerivationIntrosAt_def Let_def) using ladder_i_L'_index apply blast using intro_at_index by blast } note introsAt = this show ?thesis apply (rule_tac x="L'" in exI) apply auto defer 1 using L' ladder_stepdown_length length_L apply auto[1] using ladder_stepdown_i_0 length_L L' apply auto[1] using ladder_stepdown_last_j L' length_L apply auto[1] apply (auto simp add: LeftDerivationLadder_def) using ldl1 apply blast using is_ladder_L' apply blast using ldfix apply blast apply (auto simp add: LeftDerivationIntros_def) apply (simp add: introsAt) done qed fun ladder_shift_j :: "nat \ ladder \ ladder" where "ladder_shift_j d [] = []" | "ladder_shift_j d ((n, j, i)#L) = ((n, j - d, i)#(ladder_shift_j d L))" definition ladder_cut_prefix :: "nat \ ladder \ ladder" where "ladder_cut_prefix d L = (ladder_shift_j d L)[0 := (ladder_n L 0, ladder_j L 0 - d, ladder_i L 0 - d)]" lemma ladder_shift_j_length: "length (ladder_shift_j d L) = length L" by (induct L, auto) lemma ladder_cut_prefix_length: shows "length (ladder_cut_prefix d L) = length L" apply (simp add: ladder_cut_prefix_def) apply (simp add: ladder_shift_j_length) done lemma ladder_shift_j_cons: "ladder_shift_j d (x#L) = (fst x, fst (snd x) - d, snd(snd x))# (ladder_shift_j d L)" apply (induct L) by (cases x, simp)+ lemma deriv_j_ladder_shift_j: "index < length L \ deriv_j (ladder_shift_j d L ! index) = deriv_j (L ! index) - d" proof (induct L arbitrary: index) case Nil then show ?case by auto next case (Cons x L) show ?case apply (subst ladder_shift_j_cons) apply (cases index) using Cons by (auto simp add: deriv_j_def) qed lemma deriv_n_ladder_shift_j: "index < length L \ deriv_n (ladder_shift_j d L ! index) = deriv_n (L ! index)" proof (induct L arbitrary: index) case Nil then show ?case by auto next case (Cons x L) show ?case apply (subst ladder_shift_j_cons) apply (cases index) using Cons by (auto simp add: deriv_n_def) qed lemma deriv_ix_ladder_shift_j: "index < length L \ deriv_ix (ladder_shift_j d L ! index) = deriv_ix (L ! index)" proof (induct L arbitrary: index) case Nil then show ?case by auto next case (Cons x L) show ?case apply (subst ladder_shift_j_cons) apply (cases index) using Cons by (auto simp add: deriv_ix_def) qed lemma ladder_cut_prefix_j: assumes index_bound: "index < length L" assumes length_L: "length L > 0" shows "ladder_j (ladder_cut_prefix d L) index = ladder_j L index - d" apply (simp add: ladder_j_def ladder_cut_prefix_def) apply (cases index) apply (auto simp add: length_L) apply (subst nth_list_update_eq) apply (simp only: ladder_shift_j_length length_L) apply (simp add: deriv_j_def) apply (subst deriv_j_ladder_shift_j) using index_bound apply arith by blast lemma hd_0_subst: "length L > 0 \ hd (L [0 := x]) = x" using hd_conv_nth by (simp add: upd_conv_take_nth_drop) lemma ladder_cut_prefix_i: assumes index_bound: "index < length L" assumes length_L: "length L > 0" shows "ladder_i (ladder_cut_prefix d L) index = ladder_i L index - d" apply (simp add: ladder_i_def ladder_cut_prefix_def) apply (cases index) apply auto[1] apply (subst hd_0_subst) using length_L ladder_shift_j_length apply metis apply (auto simp add: deriv_i_def) apply (case_tac nat) apply (simp add: ladder_j_def deriv_j_def) apply auto apply (subst nth_list_update_eq) using length_L ladder_shift_j_length apply auto[1] apply simp apply (simp add: ladder_j_def) apply (subst deriv_j_ladder_shift_j) using index_bound apply arith apply simp done lemma ladder_cut_prefix_n: assumes index_bound: "index < length L" assumes length_L: "length L > 0" shows "ladder_n (ladder_cut_prefix d L) index = ladder_n L index" apply (simp add: ladder_cut_prefix_def) apply (cases index) apply (auto simp add: ladder_n_def) apply (subst nth_list_update_eq) apply (simp add: ladder_shift_j_length) using length_L apply blast apply (simp add: deriv_n_def ) apply (rule_tac deriv_n_ladder_shift_j) using index_bound by arith lemma ladder_cut_prefix_ix: assumes index_bound: "index < length L" assumes length_L: "length L > 0" shows "ladder_ix (ladder_cut_prefix d L) index = ladder_ix L index" apply (simp add: ladder_cut_prefix_def) apply (cases index) apply (auto simp add: ladder_ix_def) apply (rule_tac deriv_ix_ladder_shift_j) using index_bound by arith lemma LeftDerivationFix_derivation_ge_is_nonterminal: assumes ldfix: "LeftDerivationFix \ i D j \" assumes derivation_ge_d: "derivation_ge D d" assumes is_nonterminal: "is_nonterminal (\ ! j)" shows "(D = [] \ \ = \ \ i = j) \ (i > d \ j \ d)" proof - have "is_nonterminal (\ ! i)" using ldfix is_nonterminal by (simp add: LeftDerivationFix_def) from LeftDerivationFix_splits_at_nonterminal[OF ldfix this] obtain U a1 a2 b1 where U: "splits_at \ i a1 U a2 \ splits_at \ j b1 U a2 \ LeftDerivation a1 D b1" by blast have "D = [] \ D \ []" by auto then show ?thesis proof (induct rule: disjCases2) case 1 then have "a1 = b1" using U by auto then have i_eq_j: "i = j" using U by (metis dual_order.strict_implies_order length_take min.absorb2 splits_at_def) from 1 have "\ = \" using ldfix LeftDerivationFix_def by auto with 1 i_eq_j show ?case by blast next case 2 have "\ a1'. LeftDerives1 a1 (fst (hd D)) (snd (hd D)) a1'" using U 2 by (metis LeftDerivation.elims(2) list.sel(1)) then obtain a1' where a1': "LeftDerives1 a1 (fst (hd D)) (snd (hd D)) a1'" by blast then have "(fst (hd D)) < length a1" using Derives1_bound LeftDerives1_implies_Derives1 by blast then have fst_less_i: "(fst (hd D)) < i" using U by (simp add: leD min.absorb2 nat_le_linear splits_at_def) have d_le_fst: "d \ fst (hd D)" using derivation_ge_d 2 by (simp add: derivation_ge_def) with fst_less_i have d_less_i: "d < i" using le_less_trans by blast have "\ b1'. LeftDerives1 b1' (fst (last D)) (snd (last D)) b1" using U 2 by (metis Derive LeftDerivation_Derive_take_LeftDerives1 LeftDerivation_implies_Derivation last_conv_nth length_0_conv order_refl take_all) then obtain b1' where b1': "LeftDerives1 b1' (fst (last D)) (snd (last D)) b1" by blast then have "fst (last D) \ length b1" using Derives1_bound' LeftDerives1_implies_Derives1 by blast then have fst_le_j: "fst (last D) \ j" using U splits_at_def by auto have "d \ fst (last D)" using derivation_ge_d 2 using derivation_ge_def last_in_set by blast with fst_le_j have "d \ j" using order.trans by blast with d_less_i show ?thesis by auto qed qed lemma LeftDerivationFix_derivation_ge: assumes ldfix: "LeftDerivationFix \ i D j \" assumes derivation_ge_d: "derivation_ge D d" shows "i = j \ (i > d \ j \ d)" proof - from LeftDerivationFix_splits_at_symbol[OF ldfix] obtain U a1 a2 b1 b2 n where U: "splits_at \ i a1 U a2 \ splits_at \ j b1 U b2 \ n \ length D \ LeftDerivation a1 (take n D) b1 \ derivation_ge (drop n D) (Suc (length b1)) \ LeftDerivation a2 (derivation_shift (drop n D) (Suc (length b1)) 0) b2 \ (n = length D \ n < length D \ is_word (b1 @ [U]))" by blast have "n = 0 \ n > 0" by auto then show ?thesis proof (induct rule: disjCases2) case 1 then have "a1 = b1" using U by auto then have i_eq_j: "i = j" using U by (metis dual_order.strict_implies_order length_take min.absorb2 splits_at_def) then show ?case by blast next case 2 obtain E where E: "E = take n D" by blast have E_nonempty: "E \ []" using E 2 using U less_nat_zero_code list.size(3) take_eq_Nil by auto have "\ a1'. LeftDerives1 a1 (fst (hd E)) (snd (hd E)) a1'" using U E E_nonempty by (metis LeftDerivation.simps(2) list.exhaust list.sel(1)) then obtain a1' where a1': "LeftDerives1 a1 (fst (hd E)) (snd (hd E)) a1'" by blast then have "(fst (hd E)) < length a1" using Derives1_bound LeftDerives1_implies_Derives1 by blast then have fst_less_i: "(fst (hd E)) < i" using U by (simp add: leD min.absorb2 nat_le_linear splits_at_def) have d_le_fst: "d \ fst (hd E)" using derivation_ge_d E_nonempty E by (simp add: LeftDerivation.elims(2) U derivation_ge_def hd_conv_nth) with fst_less_i have d_less_i: "d < i" using le_less_trans by blast have "\ b1'. LeftDerives1 b1' (fst (last E)) (snd (last E)) b1" using E_nonempty U E by (metis LeftDerivation_append1 append_butlast_last_id prod.collapse) then obtain b1' where b1': "LeftDerives1 b1' (fst (last E)) (snd (last E)) b1" by blast then have "fst (last E) \ length b1" using Derives1_bound' LeftDerives1_implies_Derives1 by blast then have fst_le_j: "fst (last E) \ j" using U splits_at_def by auto have "d \ fst (last E)" using derivation_ge_d E_nonempty E using derivation_ge_d last_in_set by (metis derivation_ge_def set_take_subset subsetCE) with fst_le_j have "d \ j" using order.trans by blast with d_less_i show ?thesis by auto qed qed lemma LeftDerivationIntro_derivation_ge: assumes ldintro: "LeftDerivationIntro \ i r ix D j \" assumes i_ge_d: "i \ d" assumes derivation_ge_d: "derivation_ge D d" shows "j \ d" proof - from iffD1[OF LeftDerivationIntro_def ldintro] obtain \ where \: "LeftDerives1 \ i r \ \ ix < length (snd r) \ snd r ! ix = \ ! j \ LeftDerivationFix \ (i + ix) D j \" by blast then have "(i + ix = j) \ (i + ix > d \ j \ d)" using LeftDerivationFix_derivation_ge derivation_ge_d by blast then show ?thesis proof (induct rule: disjCases2) case 1 then show ?case using i_ge_d trans_le_add1 by blast next case 2 then show ?case by simp qed qed lemma derivation_ge_LeftDerivationLadder: assumes derivation_ge_d: "derivation_ge D d" assumes ladder: "LeftDerivationLadder \ D L \" assumes ladder_i_0: "ladder_i L 0 \ d" shows "index < length L \ ladder_i L index \ d \ ladder_j L index \ d" proof (induct index) case 0 from iffD1[OF LeftDerivationLadder_def ladder] have ldfix: "LeftDerivationFix \ (ladder_i L 0) (take (ladder_n L 0) D) (ladder_j L 0) (ladder_\ \ D L 0)" by blast have "derivation_ge (take (ladder_n L 0) D) d" using derivation_ge_d by (metis append_take_drop_id derivation_ge_append) from ladder_i_0 derivation_ge_d LeftDerivationFix_derivation_ge[OF ldfix this] show ?case by linarith next case (Suc n) have ladder_i_Suc: "ladder_i L (Suc n) \ d" apply (auto simp add: ladder_i_def) using Suc by auto from iffD1[OF LeftDerivationLadder_def ladder] have "LeftDerivationIntros \ D L" by blast then have "LeftDerivationIntrosAt \ D L (Suc n)" using Suc.prems by (metis LeftDerivationIntros_def Suc_eq_plus1_left le_add1) from iffD1[OF LeftDerivationIntrosAt_def this] show ?case using ladder_i_Suc LeftDerivationIntro_derivation_ge by (metis append_take_drop_id derivation_ge_append derivation_ge_d) qed lemma derivation_shift_append: "derivation_shift (A@B) left right = (derivation_shift A left right) @ (derivation_shift B left right)" by (induct A, simp+) lemma derivation_shift_right_left_subtract: "right \ left \ derivation_shift (derivation_shift L 0 right) left 0 = derivation_shift L 0 (right - left)" by (induct L, simp+) lemma LeftDerivationFix_cut_prefix: assumes "LeftDerivationFix (\@\) i D j \" assumes "derivation_ge D (length \)" assumes "i \ length \" assumes is_word_\: "is_word \" shows "\ \'. \ = \ @ \' \ LeftDerivationFix \ (i - length \) (derivation_shift D (length \) 0) (j - length \) \'" proof - have j_ge_d: "j \ length \" using assms(3) LeftDerivationFix_derivation_ge[OF assms(1) assms(2)] by arith obtain \' where \': "\' = drop (length \) \" by blast from iffD1[OF LeftDerivationFix_def assms(1)] obtain E F where EF: "is_sentence (\ @ \) \ is_sentence \ \ LeftDerivation (\ @ \) D \ \ i < length (\ @ \) \ j < length \ \ (\ @ \) ! i = \ ! j \ D = E @ derivation_shift F 0 (Suc j) \ LeftDerivation (take i (\ @ \)) E (take j \) \ LeftDerivation (drop (Suc i) (\ @ \)) F (drop (Suc j) \)" by blast then have "LeftDerivation (\ @ \) D \" by blast from LeftDerivation_skip_prefixword_ex[OF this is_word_\] obtain \' where \': "\ = \ @ \' \ LeftDerivation \ (derivation_shift D (length \) 0) \'" by blast have ldf1: "is_sentence \" using EF is_sentence_concat by blast have ldf2: "is_sentence \'" using EF \' is_sentence_concat by blast have ldf3: "i - length \ < length \" by (metis EF append_Nil assms(3) drop_append drop_eq_Nil not_le) have ldf4: "j - length \ < length \'" by (metis EF append_Nil j_ge_d \' drop_append drop_eq_Nil not_le) have ldf5: "\ ! (i - length \) = \' ! (j - length \)" by (metis \' EF assms(3) j_ge_d leD nth_append) have D_split: "D = E @ derivation_shift F 0 (Suc j)" using EF by blast show ?thesis apply (rule_tac x=\' in exI) apply (auto simp add: \') apply (auto simp add: LeftDerivationFix_def) using ldf1 apply blast using ldf2 apply blast using \' apply blast using ldf3 apply blast using ldf4 apply blast using ldf5 apply blast apply (rule_tac x="derivation_shift E (length \) 0" in exI) apply (rule_tac x="F" in exI) apply auto apply (subst D_split) apply (simp add: derivation_shift_append) apply (subst derivation_shift_right_left_subtract) apply (simp add: j_ge_d le_Suc_eq) using j_ge_d apply (simp add: Suc_diff_le) apply (metis EF LeftDerivation_implies_Derivation LeftDerivation_skip_prefix \' append_eq_conv_conj assms(3) drop_take is_word_Derivation_derivation_ge is_word_\ take_all take_append) using EF Suc_diff_le \' assms(3) j_ge_d by auto qed lemma LeftDerives1_propagate_prefix: "LeftDerives1 (\ @ \) i r \ \ i \ length \ \ is_prefix \ \" proof - assume a1: "LeftDerives1 (\ @ \) i r \" assume a2: "length \ \ i" have f3: "take i (\ @ \) = take i \" using a1 Derives1_take LeftDerives1_implies_Derives1 by blast then have f4: "length (take i \) = i" using a1 by (metis (no_types) Derives1_bound LeftDerives1_implies_Derives1 dual_order.strict_implies_order length_take min.absorb2) have "take (length \) (take i \) = \" using f3 a2 by (simp add: append_eq_conv_conj) then show ?thesis using f4 a2 by (metis (no_types) append_Nil2 append_eq_conv_conj diff_is_0_eq' is_prefix_take take_0 take_append) qed lemma LeftDerivationIntro_cut_prefix: assumes "LeftDerivationIntro (\@\) i r ix D j \" assumes "derivation_ge D (length \)" assumes "i \ length \" assumes is_word_\: "is_word \" shows "\ \'. \ = \ @ \' \ LeftDerivationIntro \ (i - length \) r ix (derivation_shift D (length \) 0) (j - length \) \'" proof - from iffD1[OF LeftDerivationIntro_def assms(1)] obtain \ where \: "LeftDerives1 (\ @ \) i r \ \ ix < length (snd r) \ snd r ! ix = \ ! j \ LeftDerivationFix \ (i + ix) D j \" by blast have "\ \'. \ = \ @ \'" using LeftDerives1_propagate_prefix \ assms(3) by (metis append_dropped_prefix) then obtain \' where \': "\ = \ @ \'" by blast with \ have "LeftDerives1 (\ @ \) i r (\ @ \')" by simp from LeftDerives1_skip_prefix[OF assms(3) this] have \_\': "LeftDerives1 \ (i - length \) r \'" by blast have ldfix: "LeftDerivationFix (\ @ \') (i + ix) D j \" using \ \' by auto have \_le_i_plus_ix: "length \ \ i + ix" using assms(3) by arith from LeftDerivationFix_cut_prefix[OF ldfix assms(2) \_le_i_plus_ix assms(4)] obtain \' where \': "\ = \ @ \' \ LeftDerivationFix \' (i + ix - length \) (derivation_shift D (length \) 0) (j - length \) \'" by blast have same_symbol: "\ ! j = \' ! (j - length \)" by (metis LeftDerivationFix_def \ \' \_le_i_plus_ix \' leD nth_append) have \'_\': "LeftDerivationFix \' (i - length \ + ix) (derivation_shift D (length \) 0) (j - length \) \'" by (simp add: \' assms(3)) show ?thesis apply (simp add: LeftDerivationIntro_def) apply (rule_tac x=\' in exI) apply (auto simp add: \') apply (rule_tac x=\' in exI) by (auto simp add: \ \_\' same_symbol \'_\') qed lemma LeftDerivationLadder_implies_LeftDerivation_at_index: assumes "LeftDerivationLadder \ D L \" assumes "index < length L" shows "LeftDerivation \ (take (ladder_n L index) D) (ladder_\ \ D L index)" using LeftDerivationLadder_def LeftDerivation_take_derive assms(1) ladder_\_def by auto lemma LeftDerivationLadder_cut_prefix_propagate: assumes ladder: "LeftDerivationLadder (\@\) D L \" assumes is_word_\: "is_word \" assumes derivation_ge_\: "derivation_ge D (length \)" assumes ladder_i_0: "ladder_i L 0 \ length \" assumes L': "L' = ladder_cut_prefix (length \) L" assumes D': "D' = derivation_shift D (length \) 0" shows "index < length L \ LeftDerivation \ (take (ladder_n L' index) D') (ladder_\ \ D' L' index) \ ladder_\ (\@\) D L index = \@(ladder_\ \ D' L' index) \ ladder_\ (\@\) D L index = \@(ladder_\ \ D' L' index)" proof (induct index) case 0 have ladder_\: "ladder_\ (\@\) D L 0 = \@(ladder_\ \ D' L' 0)" by (simp add: ladder_\_def) have ldfix: "LeftDerivationFix (\@\) (ladder_i L 0) (take (ladder_n L 0) D) (ladder_j L 0) (ladder_\ (\@\) D L 0)" using ladder LeftDerivationLadder_def by blast have dge_take: "derivation_ge (take (ladder_n L 0) D) (length \)" using derivation_ge_\ by (metis append_take_drop_id derivation_ge_append) from LeftDerivationFix_cut_prefix[OF ldfix dge_take ladder_i_0 is_word_\] obtain \' where \': "ladder_\ (\ @ \) D L 0 = \ @ \' \ LeftDerivationFix \ (ladder_i L 0 - length \) (derivation_shift (take (ladder_n L 0) D) (length \) 0) (ladder_j L 0 - length \) \'" by blast have ladder_\: "ladder_\ (\@\) D L 0 = \@(ladder_\ \ D' L' 0)" using \' by (metis "0.prems" D' Derive L' LeftDerivationFix_def LeftDerivation_implies_Derivation ladder_\_def ladder_cut_prefix_n take_derivation_shift) have "LeftDerivation \ (take (ladder_n L' 0) D') (ladder_\ \ D' L' 0)" proof - have "LeftDerivation (\@\) (take (ladder_n L 0) D) (ladder_\ (\@\) D L 0)" using LeftDerivationLadder_implies_LeftDerivation_at_index ladder "0.prems" by blast then show ?thesis by (metis D' LeftDerivationLadder_def LeftDerivation_skip_prefix LeftDerivation_take_derive derivation_ge_\ ladder ladder_\_def) qed then show ?case using ladder_\ ladder_\ by auto next case (Suc index) have index_less_L: "index < length L" using Suc(2) by arith then have induct: "ladder_\ (\@\) D L index = \@(ladder_\ \ D' L' index)" using Suc by blast then have ladder_\: "ladder_\ (\@\) D L (Suc index) = \@(ladder_\ \ D' L' (Suc index))" by (simp add: ladder_\_def) have introsAt: "LeftDerivationIntrosAt (\@\) D L (Suc index)" using Suc(2) ladder by (metis LeftDerivationIntros_def LeftDerivationLadder_def Suc_eq_plus1_left le_add1) obtain n m e E where n: "n = ladder_n L (Suc index - 1)" and m: "m = ladder_n L (Suc index)" and e: "e = D ! n" and E: "E = drop (Suc n) (take m D)" by blast from iffD1[OF LeftDerivationIntrosAt_def introsAt] have "LeftDerivationIntro (ladder_\ (\ @ \) D L (Suc index)) (ladder_i L (Suc index)) (snd e) (ladder_ix L (Suc index)) E (ladder_j L (Suc index)) (ladder_\ (\ @ \) D L (Suc index))" using n m e E Let_def by meson then have ldintro: "LeftDerivationIntro (\@(ladder_\ \ D' L' (Suc index))) (ladder_i L (Suc index)) (snd e) (ladder_ix L (Suc index)) E (ladder_j L (Suc index)) (ladder_\ (\ @ \) D L (Suc index))" by (simp add: ladder_\) have dge_E_\: "derivation_ge E (length \)" apply (simp add: E) using derivation_ge_\ by (metis append_take_drop_id derivation_ge_append) have ladder_i_Suc: "length \ \ ladder_i L (Suc index)" using Suc.prems derivation_ge_LeftDerivationLadder derivation_ge_\ ladder ladder_i_0 by blast from LeftDerivationIntro_cut_prefix[OF ldintro dge_E_\ ladder_i_Suc is_word_\] obtain \' where \': "ladder_\ (\ @ \) D L (Suc index) = \ @ \' \ LeftDerivationIntro (ladder_\ \ D' L' (Suc index)) (ladder_i L (Suc index) - length \) (snd e) (ladder_ix L (Suc index)) (derivation_shift E (length \) 0) (ladder_j L (Suc index) - length \) \'" by blast then have "LeftDerivation (ladder_\ \ D' L' (Suc index)) ((ladder_i L (Suc index) - length \, snd e) # (derivation_shift E (length \) 0)) \'" using LeftDerivationIntro_implies_LeftDerivation by blast then have "LeftDerivation (ladder_\ \ D' L' index) ((ladder_i L (Suc index) - length \, snd e) # (derivation_shift E (length \) 0)) \'" by (auto simp add: ladder_\_def) have ld: "LeftDerivation \ (take (ladder_n L' (Suc index)) D') (ladder_\ \ D' L' (Suc index))" proof - have "LeftDerivation (\@\) (take (ladder_n L (Suc index)) D) (ladder_\ (\@\) D L (Suc index))" using LeftDerivationLadder_implies_LeftDerivation_at_index ladder Suc.prems by blast then show ?thesis by (metis D' LeftDerivationLadder_def LeftDerivation_skip_prefix LeftDerivation_take_derive derivation_ge_\ ladder ladder_\_def) qed then show ?case using \' D' Derive L' LeftDerivationIntro_def n m e E ld LeftDerivation_implies_Derivation ladder_\_def ladder_cut_prefix_n take_derivation_shift by (metis (no_types, lifting) LeftDerivationLadder_implies_LeftDerivation_at_index LeftDerivation_skip_prefixword_ex Suc.prems Suc_leI index_less_L is_word_\ ladder ladder_\ le_0_eq neq0_conv) qed theorem LeftDerivationLadder_cut_prefix: assumes ladder: "LeftDerivationLadder (\@\) D L \" assumes is_word_\: "is_word \" assumes ladder_i_0: "ladder_i L 0 \ length \" shows "\ D' L' \'. \ = \ @ \' \ LeftDerivationLadder \ D' L' \' \ D' = derivation_shift D (length \) 0 \ length L' = length L \ ladder_i L' 0 + length \ = ladder_i L 0 \ ladder_last_j L' + length \ = ladder_last_j L" proof - obtain D' where D': "D' = derivation_shift D (length \) 0" by blast obtain L' where L': "L' = ladder_cut_prefix (length \) L" by blast obtain \' where \': "\' = drop (length \) \" by blast have ladder_last_j_upper_bound: "ladder_last_j L < length \" using ladder using ladder_last_j_bound by blast have derivation_ge_\: "derivation_ge D (length \)" using is_word_\ LeftDerivationLadder_def LeftDerivation_implies_Derivation is_word_Derivation_derivation_ge ladder by blast note derivation_ge_ladder = derivation_ge_LeftDerivationLadder[OF derivation_ge_\ ladder ladder_i_0] have ladder_last_j_lower_bound: "ladder_last_j L \ length \" using LeftDerivationLadder_def derivation_ge_ladder is_ladder_def ladder ladder_last_j_def by auto from ladder_last_j_upper_bound ladder_last_j_lower_bound have \_less_\: "length \ < length \" by arith then have \_def: "\ = \ @ \'" by (metis LeftDerivation.simps(1) LeftDerivationLadder_def LeftDerivation_ge_take \' append_eq_conv_conj derivation_ge_\ ladder) have length_L_nonzero: "length L \ 0" using LeftDerivationLadder_def is_ladder_def ladder by auto have ladder_i_L'_thm: "\ index. index < length L \ ladder_i L' index + length \ = ladder_i L index" apply (simp add: L') apply (subst ladder_cut_prefix_i) apply simp using length_L_nonzero apply blast using derivation_ge_ladder by auto have ladder_j_L'_thm: "\ index. index < length L \ ladder_j L' index + length \ = ladder_j L index" apply (simp add: L') apply (subst ladder_cut_prefix_j) using LeftDerivationLadder_def is_ladder_def ladder apply blast using LeftDerivationLadder_def is_ladder_def ladder apply blast using derivation_ge_ladder by auto have length_L': "length L' = length L" using L' ladder_cut_prefix_length by simp have \_\': "LeftDerivation \ D' \'" using D' LeftDerivationLadder_def LeftDerivation_skip_prefix \' derivation_ge_\ ladder by blast have length_D': "length D' = length D" by (simp add: D') have is_ladder_D_L: "is_ladder D L" using LeftDerivationLadder_def ladder by blast { fix u :: nat assume u_bound_L': "u < length L'" have u_bound_L: "u < length L" using length_L' using u_bound_L' by simp have "ladder_n L' u \ length D'" apply (simp add: length_D' L') apply (subst ladder_cut_prefix_n) apply (simp add: u_bound_L) using length_L_nonzero apply arith using u_bound_L is_ladder_D_L by (simp add: is_ladder_def) } note is_ladder_1 = this { fix u :: nat fix v :: nat assume u_less_v: "u < v" assume v_bound_L': "v < length L'" then have v_bound_L: "v < length L" by (simp add: length_L') with u_less_v have u_bound_L: "u < length L" by arith have "ladder_n L' u < ladder_n L' v" apply (simp add: L') apply (subst ladder_cut_prefix_n) using u_bound_L apply blast using length_L_nonzero apply blast apply (subst ladder_cut_prefix_n) using v_bound_L apply blast using length_L_nonzero apply blast using u_less_v v_bound_L is_ladder_D_L by (simp add: is_ladder_def) } note is_ladder_2 = this have is_ladder_3: "ladder_last_n L' = length D'" apply (simp add: length_D' ladder_last_n_def L') apply (subst ladder_cut_prefix_n) apply (simp add: ladder_cut_prefix_length) using length_L_nonzero apply auto[1] using length_L_nonzero apply blast apply (simp add: ladder_cut_prefix_length) using is_ladder_D_L by (simp add: is_ladder_def ladder_last_n_def) have is_ladder_4: "LeftDerivationFix \ (ladder_i L' 0) (take (ladder_n L' 0) D') (ladder_j L' 0) (ladder_\ \ D' L' 0)" proof - have ldfix: "LeftDerivationFix (\@\) (ladder_i L 0) (take (ladder_n L 0) D) (ladder_j L 0) (ladder_\ (\@\) D L 0)" using ladder LeftDerivationLadder_def by blast have dge: "derivation_ge (take (ladder_n L 0) D) (length \)" using derivation_ge_\ by (metis append_take_drop_id derivation_ge_append) from LeftDerivationFix_cut_prefix[OF ldfix dge ladder_i_0 is_word_\] obtain \' where \': "ladder_\ (\ @ \) D L 0 = \ @ \' \ LeftDerivationFix \ (ladder_i L 0 - length \) (derivation_shift (take (ladder_n L 0) D) (length \) 0) (ladder_j L 0 - length \) \'" by blast then show ?thesis using LeftDerivationLadder_cut_prefix_propagate D' L' append_eq_conv_conj derivation_ge_\ is_word_\ ladder ladder_cut_prefix_i ladder_cut_prefix_j ladder_cut_prefix_n ladder_i_0 length_0_conv length_L_nonzero length_greater_0_conv take_derivation_shift by auto qed { fix index :: nat assume index_lower_bound: "Suc 0 \ index" assume index_upper_bound: "index < length L'" have introsAt: "LeftDerivationIntrosAt (\@\) D L index" by (metis LeftDerivationIntros_def LeftDerivationLadder_def One_nat_def index_lower_bound index_upper_bound ladder length_L') then have ladder_i_L: "ladder_i L index = fst (D ! ladder_n L (index - Suc 0))" by (metis LeftDerivationIntrosAt_def One_nat_def \LeftDerivationIntrosAt (\ @ \) D L index\) have ladder_i_L'_as_L: "ladder_i L' index = ladder_i L index - (length \)" using ladder_cut_prefix_i L' index_upper_bound is_ladder_D_L is_ladder_not_empty length_L' length_greater_0_conv by auto have length_L_gr_0: "length L > 0" using length_L' length_L_nonzero by arith have index_Suc_upper_bound_L: "index - Suc 0 < length L" using index_upper_bound length_L' by arith have "fst (D' ! ladder_n L' (index - Suc 0)) = fst (D ! ladder_n L (index - Suc 0)) - (length \)" apply (subst D', subst L') apply (subst ladder_cut_prefix_n[OF index_Suc_upper_bound_L length_L_gr_0]) apply (simp add: derivation_shift_def) using index_lower_bound index_upper_bound is_ladder_D_L ladder_n_minus_1_bound length_L' by auto then have ladder_i_L': "ladder_i L' index = fst (D' ! ladder_n L' (index - Suc 0))" using ladder_i_L ladder_i_L'_as_L by auto have "LeftDerivationIntro (ladder_\ \ D' L' index) (ladder_i L' index) (snd (D' ! ladder_n L' (index - Suc 0))) (ladder_ix L' index) (drop (Suc (ladder_n L' (index - Suc 0))) (take (ladder_n L' index) D')) (ladder_j L' index) (ladder_\ \ D' L' index)" proof - have "LeftDerivationIntro (ladder_\ (\@\) D L index) (ladder_i L index) (snd (D ! ladder_n L (index - Suc 0))) (ladder_ix L index) (drop (Suc (ladder_n L (index - Suc 0))) (take (ladder_n L index) D)) (ladder_j L index) (ladder_\ (\@\) D L index)" using introsAt by (metis LeftDerivationIntrosAt_def One_nat_def) then have ldintro: "LeftDerivationIntro (\@(ladder_\ \ D' L' index)) (ladder_i L index) (snd (D ! ladder_n L (index - Suc 0))) (ladder_ix L index) (drop (Suc (ladder_n L (index - Suc 0))) (take (ladder_n L index) D)) (ladder_j L index) (ladder_\ (\@\) D L index)" using D' L' LeftDerivationLadder_cut_prefix_propagate derivation_ge_\ index_upper_bound is_word_\ ladder ladder_i_0 length_L' by auto have dge: "derivation_ge (drop (Suc (ladder_n L (index - Suc 0))) (take (ladder_n L index) D)) (length \)" using derivation_ge_\ by (metis append_take_drop_id derivation_ge_append) have \_le_i_L: "length \ \ ladder_i L index" using derivation_ge_ladder index_upper_bound length_L' by auto from LeftDerivationIntro_cut_prefix[OF ldintro dge \_le_i_L is_word_\] obtain \' where \': "ladder_\ (\ @ \) D L index = \ @ \' \ LeftDerivationIntro (ladder_\ \ D' L' index) (ladder_i L index - length \) (snd (D ! ladder_n L (index - Suc 0))) (ladder_ix L index) (derivation_shift (drop (Suc (ladder_n L (index - Suc 0))) (take (ladder_n L index) D)) (length \) 0) (ladder_j L index - length \) \'" by blast have h1: "ladder_i L' index = ladder_i L index - length \" using L' ladder_cut_prefix_i ladder_i_L'_as_L by blast have h2: "(snd (D' ! ladder_n L' (index - Suc 0))) = (snd (D ! ladder_n L (index - Suc 0)))" apply (subst L', subst ladder_cut_prefix_n) apply (simp add: index_Suc_upper_bound_L) using length_L_gr_0 apply auto[1] apply (subst D') apply (simp add: derivation_shift_def) using index_lower_bound index_upper_bound is_ladder_D_L ladder_n_minus_1_bound length_L' by auto have h3: "ladder_ix L' index = ladder_ix L index" using ladder_cut_prefix_ix L' index_upper_bound length_L' length_L_gr_0 by auto have h4: "(drop (Suc (ladder_n L' (index - Suc 0))) (take (ladder_n L' index) D')) = (derivation_shift (drop (Suc (ladder_n L (index - Suc 0))) (take (ladder_n L index) D)) (length \) 0)" apply (subst D') apply (subst L')+ apply (subst ladder_cut_prefix_n, simp add: index_Suc_upper_bound_L) using length_L_gr_0 apply blast apply (subst ladder_cut_prefix_n) using index_upper_bound length_L' apply arith using length_L_gr_0 apply blast apply (simp add: derivation_shift_def) by (simp add: drop_map take_map) have h5: "ladder_j L' index = ladder_j L index - length \" using ladder_cut_prefix_j L' index_upper_bound length_L' length_L_gr_0 by auto have h6: "ladder_\ \ D' L' index = \'" using D' L' LeftDerivationLadder_cut_prefix_propagate \' derivation_ge_\ index_upper_bound is_word_\ ladder ladder_i_0 length_L' by auto show ?thesis using h1 h2 h3 h4 h5 h6 \' by simp qed then have "LeftDerivationIntrosAt \ D' L' index" apply (auto simp add: LeftDerivationIntrosAt_def Let_def) using ladder_i_L' by blast } note is_ladder_5 = this show ?thesis apply (rule_tac x="D'" in exI) apply (rule_tac x="L'" in exI) apply (rule_tac x="\'" in exI) apply auto using \_def apply blast defer 1 using D' apply blast using L' ladder_cut_prefix_length apply auto[1] apply (subst ladder_i_L'_thm) using LeftDerivationLadder_def is_ladder_def ladder apply blast apply simp apply (simp add: ladder_last_j_def) apply (subst ladder_j_L'_thm) apply (simp add: length_L') using length_L_nonzero apply arith apply (simp add: length_L') apply (auto simp add: LeftDerivationLadder_def) using \_\' apply blast apply (auto simp add: is_ladder_def) using length_L_nonzero length_L' apply auto[1] using is_ladder_1 apply blast using is_ladder_2 apply blast using is_ladder_3 apply blast using is_ladder_4 apply blast by (auto simp add: LeftDerivationIntros_def is_ladder_5) qed end end