Nuprl Lemma : opp-side_half-plane-angle-congruence-lemma

e:EuclideanPlane. ∀b,p,b',p',a,c,a',c',d:Point.
  ((a leftof bp ∧ leftof pb)
   (a' leftof b'p' ∧ c' leftof p'b')
   ((abp ≅a a'b'p' ∧ pbc ≅a p'b'c') ∧ (a-d-c ∧ out(b dp)) ∧ b ≠ d)
   (∃a'',c'',d'':Point
       ((ba ≅ b'a'' ∧ (bc ≅ b'c'' ∧ bd ≅ b'd'') ∧ ad ≅ a''d'' ∧ dc ≅ d''c'' ∧ out(b' d''p')) ∧ a''_d''_c'')))


Proof




Definitions occuring in Statement :  geo-out: out(p ab) geo-cong-angle: abc ≅a xyz euclidean-plane: EuclideanPlane geo-strict-between: a-b-c geo-left: leftof bc geo-congruent: ab ≅ cd geo-between: a_b_c geo-sep: a ≠ b geo-point: Point all: x:A. B[x] exists: x:A. B[x] implies:  Q and: P ∧ Q
Definitions unfolded in proof :  all: x:A. B[x] implies:  Q and: P ∧ Q member: t ∈ T prop: uall: [x:A]. B[x] basic-geometry: BasicGeometry subtype_rel: A ⊆B guard: {T} uimplies: supposing a exists: x:A. B[x] cand: c∧ B uiff: uiff(P;Q) so_lambda: λ2x.t[x] so_apply: x[s] geo-strict-between: a-b-c iff: ⇐⇒ Q basic-geometry-: BasicGeometry- oriented-plane: OrientedPlane
Lemmas referenced :  geo-cong-angle_wf geo-strict-between_wf euclidean-plane-structure-subtype euclidean-plane-subtype subtype_rel_transitivity euclidean-plane_wf euclidean-plane-structure_wf geo-primitives_wf geo-out_wf geo-sep_wf geo-left_wf geo-point_wf left-implies-sep geo-proper-extend-exists geo-sep-sym geo-strict-between-sep3 geo-congruent-iff-length geo-congruent_wf exists_wf geo-out-iff-between1 geo-between-symmetry euclidean-plane-axioms geo-strict-between-implies-between geo-out_inversion geo-between_wf geo-out_weakening geo-eq_weakening geo-sas2 out-preserves-angle-cong_1 geo-congruent-symmetry geo-congruent-sep geo-strict-between-sep2 geo-five-segment geo-congruent-comm Euclid-Prop7 left-between-implies-right1 geo-left-out geo-left-out-2 geo-left-out-1 geo-left-out-3 geo-congruent_functionality geo-sep_functionality geo-between_functionality
Rules used in proof :  sqequalSubstitution sqequalTransitivity computationStep sqequalReflexivity lambdaFormation sqequalHypSubstitution productElimination thin cut productEquality introduction extract_by_obid isectElimination sqequalRule hypothesisEquality hypothesis applyEquality instantiate independent_isectElimination because_Cache dependent_functionElimination independent_functionElimination rename dependent_pairFormation independent_pairFormation equalitySymmetry lambdaEquality equalityTransitivity dependent_set_memberEquality promote_hyp

Latex:
\mforall{}e:EuclideanPlane.  \mforall{}b,p,b',p',a,c,a',c',d:Point.
    ((a  leftof  bp  \mwedge{}  c  leftof  pb)
    {}\mRightarrow{}  (a'  leftof  b'p'  \mwedge{}  c'  leftof  p'b')
    {}\mRightarrow{}  ((abp  \mcong{}\msuba{}  a'b'p'  \mwedge{}  pbc  \mcong{}\msuba{}  p'b'c')  \mwedge{}  (a-d-c  \mwedge{}  out(b  dp))  \mwedge{}  b  \mneq{}  d)
    {}\mRightarrow{}  (\mexists{}a'',c'',d'':Point
              ((ba  \mcong{}  b'a''  \mwedge{}  (bc  \mcong{}  b'c''  \mwedge{}  bd  \mcong{}  b'd'')  \mwedge{}  ad  \mcong{}  a''d''  \mwedge{}  dc  \mcong{}  d''c''  \mwedge{}  out(b'  d''p'))
              \mwedge{}  a''\_d''\_c'')))



Date html generated: 2018_05_22-PM-00_20_03
Last ObjectModification: 2018_04_21-PM-10_36_54

Theory : euclidean!plane!geometry


Home Index