Nuprl Lemma : pW-rec_wf

[P:Type]. ∀[A:P ⟶ Type]. ∀[B:p:P ⟶ A[p] ⟶ Type]. ∀[C:p:P ⟶ a:A[p] ⟶ B[p;a] ⟶ P]. ∀[Q:par:P ⟶ (pW par) ⟶ ℙ].
[ind:par:P
      ⟶ a:A[par]
      ⟶ f:(b:B[par;a] ⟶ (pW C[par;a;b]))
      ⟶ (b:B[par;a] ⟶ Q[C[par;a;b];f b])
      ⟶ Q[par;pW-sup(a;f)]]. ∀[par:P]. ∀[w:pW par].
  (let ind(p,a,f,G) ind[p;a;f;G] in 
   letrec F(p,w) let a,f=w in 
                   ind(p,a,f,λb.F(C[p;a;b],f(b)) in 
   F(par;w) ∈ Q[par;w])


Proof




Definitions occuring in Statement :  pW-rec: pW-rec pW-sup: pW-sup(a;f) param-W: pW uall: [x:A]. B[x] prop: so_apply: x[s1;s2;s3;s4] so_apply: x[s1;s2;s3] so_apply: x[s1;s2] so_apply: x[s] member: t ∈ T apply: a function: x:A ⟶ B[x] universe: Type
Definitions unfolded in proof :  uall: [x:A]. B[x] member: t ∈ T subtype_rel: A ⊆B decidable: Dec(P) all: x:A. B[x] nat: so_apply: x[s] so_apply: x[s1;s2] so_apply: x[s1;s2;s3] prop: so_lambda: λ2x.t[x] so_lambda: λ2y.t[x; y] so_lambda: so_lambda(x,y,z.t[x; y; z]) param-W: pW pcw-partial: pcw-partial(path;n) implies:  Q squash: T pcw-path: Path or: P ∨ Q iff: ⇐⇒ Q and: P ∧ Q not: ¬A rev_implies:  Q false: False uiff: uiff(P;Q) uimplies: supposing a sq_stable: SqStable(P) subtract: m top: Top le: A ≤ B less_than': less_than'(a;b) true: True 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 lelt: i ≤ j < k int_seg: {i..j-} pcw-pp-barred: Barred(pp) spreadn: spread3 isr: isr(x) exposed-it: exposed-it ext-family: F ≡ G ext-eq: A ≡ B pi1: fst(t) pW-rec: pW-rec pW-sup: pW-sup(a;f) so_apply: x[s1;s2;s3;s4] pcw-steprel: StepRel(s1;s2) pcw-step-agree: StepAgree(s;p1;w) cand: c∧ B
Lemmas referenced :  pcw-step_wf param-W_wf param-W-rel_wf pcw-pp-barred_wf0 decidable__pcw-pp-barred int_seg_wf istype-universe subtype_rel_self pW-sup_wf decidable__le istype-false not-le-2 sq_stable__le condition-implies-le minus-add istype-void 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 nat_wf false_wf true_wf equal_wf lelt_wf le-add-cancel-alt zero-mul add-mul-special decidable__lt minus-minus less-iff-le subtract_wf top_wf squash_wf equal-wf-base member_wf param-co-W-ext param-co-W_wf subtype_rel-equal param-W-ext pcw-pp-lemma equal-implies-member-param-W subtype_rel_weakening unit_wf2 iff_weakening_equal set_subtype_base int_subtype_base decidable__int_equal not-equal-2 minus-zero subtype_rel_function int_seg_subtype
Rules used in proof :  cut introduction extract_by_obid sqequalSubstitution sqequalTransitivity computationStep sqequalReflexivity Error :isect_memberFormation_alt,  hypothesis sqequalHypSubstitution isectElimination thin hypothesisEquality because_Cache strong_bar_Induction equalityTransitivity equalitySymmetry applyEquality sqequalRule dependent_functionElimination Error :dependent_pairEquality_alt,  Error :functionIsType,  Error :universeIsType,  natural_numberEquality setElimination rename axiomEquality Error :isect_memberEquality_alt,  instantiate universeEquality Error :lambdaEquality_alt,  Error :inhabitedIsType,  independent_functionElimination imageElimination imageMemberEquality baseClosed Error :dependent_set_memberEquality_alt,  Error :lambdaFormation_alt,  addEquality unionElimination independent_pairFormation voidElimination productElimination independent_isectElimination minusEquality equalityElimination lessCases axiomSqEquality Error :productIsType,  Error :equalityIsType1,  Error :dependent_pairFormation_alt,  promote_hyp cumulativity functionExtensionality lambdaEquality lambdaFormation dependent_set_memberEquality intEquality voidEquality isect_memberEquality isect_memberFormation productEquality Error :setIsType,  int_eqReduceTrueSq hypothesis_subsumption applyLambdaEquality functionEquality Error :inlEquality_alt,  Error :unionIsType,  hyp_replacement

Latex:
\mforall{}[P:Type].  \mforall{}[A:P  {}\mrightarrow{}  Type].  \mforall{}[B:p:P  {}\mrightarrow{}  A[p]  {}\mrightarrow{}  Type].  \mforall{}[C:p:P  {}\mrightarrow{}  a:A[p]  {}\mrightarrow{}  B[p;a]  {}\mrightarrow{}  P].  \mforall{}[Q:par:P
                                                                                                                                                                                        {}\mrightarrow{}  (pW 
                                                                                                                                                                                                par)
                                                                                                                                                                                        {}\mrightarrow{}  \mBbbP{}].
\mforall{}[ind:par:P
            {}\mrightarrow{}  a:A[par]
            {}\mrightarrow{}  f:(b:B[par;a]  {}\mrightarrow{}  (pW  C[par;a;b]))
            {}\mrightarrow{}  (b:B[par;a]  {}\mrightarrow{}  Q[C[par;a;b];f  b])
            {}\mrightarrow{}  Q[par;pW-sup(a;f)]].  \mforall{}[par:P].  \mforall{}[w:pW  par].
    (let  ind(p,a,f,G)  =  ind[p;a;f;G]  in 
      letrec  F(p,w)  =  let  a,f=w  in 
                                      ind(p,a,f,\mlambda{}b.F(C[p;a;b],f(b))  in 
      F(par;w)  \mmember{}  Q[par;w])



Date html generated: 2019_06_20-PM-00_36_17
Last ObjectModification: 2018_10_06-AM-11_20_30

Theory : co-recursion


Home Index