Nuprl Lemma : cw-pp-lemma

A:Type. ∀B:A ⟶ Type. ∀w:W(A;a.B[a]). ∀n:ℕ. ∀m:ℕn. ∀ss:ℕn ⟶ cw-step(A;a.B[a]).
  ((∀x:ℕn. (W-rel(A;a.B[a];w) ss (ss x)))  (fst(snd((ss m))) ∈ W(A;a.B[a])))


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]) int_seg: {i..j-} nat: so_apply: x[s] pi1: fst(t) pi2: snd(t) all: x:A. B[x] implies:  Q member: t ∈ T apply: a function: x:A ⟶ B[x] natural_number: $n universe: Type
Definitions unfolded in proof :  all: x:A. B[x] implies:  Q member: t ∈ T W: W(A;a.B[a]) so_lambda: λ2x.t[x] so_apply: x[s] 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] subtype_rel: A ⊆B cw-step: cw-step(A;a.B[a]) uall: [x:A]. B[x] prop: nat: uimplies: supposing a le: A ≤ B and: P ∧ Q less_than': less_than'(a;b) false: False not: ¬A int_seg: {i..j-} sq_stable: SqStable(P) lelt: i ≤ j < k guard: {T} squash: T W-rel: W-rel(A;a.B[a];w) pcw-step: pcw-step(P;p.A[p];p,a.B[p; a];p,a,b.C[p; a; b]) pi1: fst(t)
Lemmas referenced :  pcw-pp-lemma unit_wf2 it_wf cw-step_wf all_wf int_seg_wf W-rel_wf int_seg_subtype_nat false_wf subtype_rel_dep_function int_seg_subtype sq_stable__le le_weakening2 subtype_rel_self nat_wf W_wf subtype_rel-equal param-W_wf equal_wf equal-unit
Rules used in proof :  sqequalSubstitution sqequalTransitivity computationStep sqequalReflexivity lambdaFormation cut introduction extract_by_obid sqequalHypSubstitution dependent_functionElimination thin hypothesis sqequalRule lambdaEquality cumulativity hypothesisEquality applyEquality functionExtensionality isectElimination because_Cache independent_functionElimination natural_numberEquality setElimination rename independent_isectElimination independent_pairFormation productElimination imageMemberEquality baseClosed imageElimination functionEquality universeEquality equalityTransitivity equalitySymmetry

Latex:
\mforall{}A:Type.  \mforall{}B:A  {}\mrightarrow{}  Type.  \mforall{}w:W(A;a.B[a]).  \mforall{}n:\mBbbN{}.  \mforall{}m:\mBbbN{}n.  \mforall{}ss:\mBbbN{}n  {}\mrightarrow{}  cw-step(A;a.B[a]).
    ((\mforall{}x:\mBbbN{}n.  (W-rel(A;a.B[a];w)  x  ss  (ss  x)))  {}\mRightarrow{}  (fst(snd((ss  m)))  \mmember{}  W(A;a.B[a])))



Date html generated: 2017_04_14-AM-07_43_30
Last ObjectModification: 2017_02_27-PM-03_14_11

Theory : co-recursion


Home Index