Nuprl Lemma : Euclid-Prop16-construction-lemma

g:EuclideanPlane. ∀a,b,c:Point.
  (a bc  (∃x:{x:Point| b=x=c} . ∃y:{y:Point| a=x=y} (abc ≅a ycb ∧ (b ≠ x ∧ x ≠ c) ∧ a ≠ x ∧ x ≠ y)))


Proof




Definitions occuring in Statement :  geo-cong-angle: abc ≅a xyz euclidean-plane: EuclideanPlane geo-lsep: bc geo-midpoint: a=m=b geo-sep: a ≠ b geo-point: Point all: x:A. B[x] exists: x:A. B[x] implies:  Q and: P ∧ Q set: {x:A| B[x]} 
Definitions unfolded in proof :  all: x:A. B[x] implies:  Q member: t ∈ T guard: {T} and: P ∧ Q uall: [x:A]. B[x] subtype_rel: A ⊆B prop: sq_exists: x:A [B[x]] euclidean-plane: EuclideanPlane sq_stable: SqStable(P) squash: T exists: x:A. B[x] cand: c∧ B basic-geometry: BasicGeometry uimplies: supposing a geo-midpoint: a=m=b geo-colinear-set: geo-colinear-set(e; L) l_all: (∀x∈L.P[x]) top: Top int_seg: {i..j-} lelt: i ≤ j < k le: A ≤ B less_than': less_than'(a;b) false: False not: ¬A less_than: a < b true: True select: L[n] cons: [a b] subtract: m basic-geometry-: BasicGeometry- uiff: uiff(P;Q) geo-strict-between: a-b-c geo-cong-angle: abc ≅a xyz
Lemmas referenced :  Euclid-midpoint lsep-implies-sep geo-sep_wf sq_stable__midpoint geo-midpoint_wf colinear-lsep lsep-all-sym geo-sep-sym midpoint-sep geo-colinear-permute geo-colinear-is-colinear-set geo-between-implies-colinear length_of_cons_lemma istype-void length_of_nil_lemma istype-false istype-le istype-less_than geo-cong-angle_wf geo-lsep_wf euclidean-plane-structure-subtype euclidean-plane-subtype subtype_rel_transitivity euclidean-plane_wf euclidean-plane-structure_wf geo-primitives_wf geo-point_wf colinear-implies-midpoint geo-strict-between-sep1 geo-strict-between-implies-colinear geo-congruent-iff-length geo-length-flip geo-strict-between-sep3 geo-proper-extend-exists geo-strict-between-implies-between geo-between-symmetry euclidean-plane-axioms geo-strict-between-sep2 geo-between-trivial vertical-angles-congruent geo-between_wf geo-congruent_wf geo-sas2 geo-between-out geo-out_weakening geo-eq_weakening out-preserves-angle-cong_1
Rules used in proof :  sqequalSubstitution sqequalTransitivity computationStep sqequalReflexivity lambdaFormation_alt cut introduction extract_by_obid sqequalHypSubstitution dependent_functionElimination thin hypothesisEquality because_Cache independent_functionElimination hypothesis productElimination dependent_set_memberEquality_alt universeIsType isectElimination applyEquality sqequalRule setElimination rename imageMemberEquality baseClosed imageElimination dependent_pairFormation_alt independent_isectElimination isect_memberEquality_alt voidElimination natural_numberEquality independent_pairFormation productIsType setIsType instantiate inhabitedIsType equalityTransitivity equalitySymmetry

Latex:
\mforall{}g:EuclideanPlane.  \mforall{}a,b,c:Point.
    (a  \#  bc
    {}\mRightarrow{}  (\mexists{}x:\{x:Point|  b=x=c\}  .  \mexists{}y:\{y:Point|  a=x=y\}  .  (abc  \mcong{}\msuba{}  ycb  \mwedge{}  (b  \mneq{}  x  \mwedge{}  x  \mneq{}  c)  \mwedge{}  a  \mneq{}  x  \mwedge{}  x  \mneq{}  y)))



Date html generated: 2019_10_16-PM-02_14_38
Last ObjectModification: 2018_12_03-PM-09_58_16

Theory : euclidean!plane!geometry


Home Index