Nuprl Lemma : W-path-lemma

A:Type. ∀B:A ⟶ Type. ∀x:W(A;a.B[a]). ∀alpha:ℕ ⟶ cw-step(A;a.B[a]).
  ((∀n:ℕ(W-rel(A;a.B[a];x) alpha (alpha n)))  (alpha ∈ Path))


Proof




Definitions occuring in Statement :  W-rel: W-rel(A;a.B[a];w) W: W(A;a.B[a]) cw-step: cw-step(A;a.B[a]) pcw-path: Path nat: it: so_apply: x[s] all: x:A. B[x] implies:  Q unit: Unit member: t ∈ T apply: a function: x:A ⟶ B[x] universe: Type
Definitions unfolded in proof :  all: x:A. B[x] implies:  Q member: t ∈ T uall: [x:A]. B[x] so_lambda: λ2x.t[x] so_apply: x[s] subtype_rel: A ⊆B nat: uimplies: supposing a le: A ≤ B and: P ∧ Q less_than': less_than'(a;b) false: False not: ¬A pcw-path: Path cw-step: cw-step(A;a.B[a]) so_lambda: λ2y.t[x; y] so_apply: x[s1;s2] so_lambda: so_lambda(x,y,z.t[x; y; z]) so_apply: x[s1;s2;s3] decidable: Dec(P) or: P ∨ Q iff: ⇐⇒ Q rev_implies:  Q uiff: uiff(P;Q) sq_stable: SqStable(P) squash: T subtract: m top: Top true: True prop: W-rel: W-rel(A;a.B[a];w) param-W-rel: param-W-rel(P;p.A[p];p,a.B[p; a];p,a,b.C[p; a; b];par;w) bool: 𝔹 unit: Unit it: btrue: tt less_than: a < b pcw-step: pcw-step(P;p.A[p];p,a.B[p; a];p,a,b.C[p; a; b]) pi2: snd(t) isl: isl(x) guard: {T} bfalse: ff exists: x:A. B[x] sq_type: SQType(T) bnot: ¬bb ifthenelse: if then else fi  assert: b W: W(A;a.B[a])
Lemmas referenced :  nat_wf W-rel_wf istype-universe subtype_rel_function cw-step_wf int_seg_wf int_seg_subtype_nat istype-void subtype_rel_self W_wf pcw-step_wf unit_wf2 it_wf decidable__le istype-false not-le-2 sq_stable__le condition-implies-le minus-add istype-int minus-one-mul zero-add minus-one-mul-top add-associates add-swap add-commutes add_functionality_wrt_le add-zero le-add-cancel le_wf add-subtract-cancel lt_int_wf eqtt_to_assert assert_of_lt_int istype-top assert_wf btrue_wf bfalse_wf pcw-steprel_wf le_weakening2 eqff_to_assert bool_cases_sqequal subtype_base_sq bool_wf bool_subtype_base assert-bnot iff_weakening_uiff less_than_wf not-lt-2 pcw-step-agree_wf subtype_rel-equal param-W_wf param-co-W_wf top_wf
Rules used in proof :  sqequalSubstitution sqequalTransitivity computationStep sqequalReflexivity Error :lambdaFormation_alt,  cut sqequalHypSubstitution hypothesis sqequalRule Error :functionIsType,  Error :universeIsType,  introduction extract_by_obid applyEquality isectElimination thin hypothesisEquality Error :lambdaEquality_alt,  natural_numberEquality setElimination rename because_Cache independent_isectElimination independent_pairFormation Error :inhabitedIsType,  universeEquality Error :dependent_set_memberEquality_alt,  functionExtensionality dependent_functionElimination addEquality unionElimination voidElimination productElimination independent_functionElimination imageMemberEquality baseClosed imageElimination Error :isect_memberEquality_alt,  minusEquality equalityElimination equalityTransitivity equalitySymmetry lessCases Error :isect_memberFormation_alt,  axiomSqEquality Error :productIsType,  Error :equalityIsType1,  Error :dependent_pairFormation_alt,  promote_hyp instantiate cumulativity

Latex:
\mforall{}A:Type.  \mforall{}B:A  {}\mrightarrow{}  Type.  \mforall{}x:W(A;a.B[a]).  \mforall{}alpha:\mBbbN{}  {}\mrightarrow{}  cw-step(A;a.B[a]).
    ((\mforall{}n:\mBbbN{}.  (W-rel(A;a.B[a];x)  n  alpha  (alpha  n)))  {}\mRightarrow{}  (alpha  \mmember{}  Path))



Date html generated: 2019_06_20-PM-00_36_24
Last ObjectModification: 2018_10_06-AM-11_20_32

Theory : co-recursion


Home Index