Nuprl Lemma : cbva_seq-list-case1

[F,G1,G2,L1,L2,a:Top]. ∀[m1,m2:ℕ].
  (cbva_seq(λn.if (n) < (m1)
                  then L1 n
                  else mk_lambdas_fun(λg.if m1=m2
                                         then mk_lambdas_fun(λg1.((G1 g) (G2 g1));m2)
                                         else (L2 (n m1));m1); F; (m1 m2) 1) 
  cbva_seq(λn.if (n) < (m1)
                   then L1 n
                   else if (n) < (m1 m2)
                           then mk_lambdas(L2 (n m1);m1)
                           else mk_lambdas_fun(λg1.mk_lambdas_fun(λg2.((G1 g1) (G2 g2));m2);m1); F; (m1 m2) 1))


Proof




Definitions occuring in Statement :  mk_lambdas: mk_lambdas(F;m) mk_lambdas_fun: mk_lambdas_fun(F;m) cbva_seq: cbva_seq(L; F; m) nat: uall: [x:A]. B[x] top: Top less: if (a) < (b)  then c  else d int_eq: if a=b then else d apply: a lambda: λx.A[x] subtract: m add: m natural_number: $n sqequal: t
Definitions unfolded in proof :  cbva_seq: cbva_seq(L; F; m) mk_applies: mk_applies(F;G;m) all: x:A. B[x] member: t ∈ T top: Top nat: uall: [x:A]. B[x] ge: i ≥  decidable: Dec(P) or: P ∨ Q uimplies: supposing a not: ¬A implies:  Q satisfiable_int_formula: satisfiable_int_formula(fmla) exists: x:A. B[x] false: False and: P ∧ Q prop: int_seg: {i..j-} lelt: i ≤ j < k le: A ≤ B less_than': less_than'(a;b) sq_type: SQType(T) guard: {T} bool: 𝔹 unit: Unit it: btrue: tt uiff: uiff(P;Q) ifthenelse: if then else fi  bfalse: ff subtype_rel: A ⊆B so_lambda: λ2x.t[x] so_apply: x[s] bnot: ¬bb assert: b rev_implies:  Q iff: ⇐⇒ Q lt_int: i <j so_lambda: so_lambda(x,y,z,w.t[x; y; z; w]) so_apply: x[s1;s2;s3;s4] has-value: (a)↓ less_than: a < b true: True squash: T so_apply: x[s1;s2]
Lemmas referenced :  primrec0_lemma istype-void nat_properties decidable__le full-omega-unsat intformand_wf intformnot_wf intformle_wf itermConstant_wf itermAdd_wf itermVar_wf istype-int int_formula_prop_and_lemma int_formula_prop_not_lemma int_formula_prop_le_lemma int_term_value_constant_lemma int_term_value_add_lemma int_term_value_var_lemma int_formula_prop_wf le_wf istype-false decidable__lt intformless_wf int_formula_prop_less_lemma less_than_wf add-subtract-cancel subtype_base_sq int_subtype_base decidable__equal_int intformeq_wf itermSubtract_wf int_formula_prop_eq_lemma int_term_value_subtract_lemma lt_int_wf eqtt_to_assert assert_of_lt_int eqff_to_assert set_subtype_base bool_cases_sqequal bool_wf bool_subtype_base assert-bnot iff_weakening_uiff assert_wf nat_wf istype-top callbyvalueall_seq-decomp-last less_as_ite callbyvalueall_seq-fun2 lifting-strict-less strict4-decide has-value_wf_base is-exception_wf callbyvalueall_seq-fun4 mk_lambdas_as_lambdas_fun
Rules used in proof :  sqequalSubstitution sqequalTransitivity computationStep sqequalReflexivity sqequalRule cut introduction extract_by_obid sqequalHypSubstitution dependent_functionElimination thin isect_memberEquality_alt voidElimination hypothesis hypothesisEquality dependent_set_memberEquality_alt addEquality setElimination rename natural_numberEquality isectElimination unionElimination independent_isectElimination approximateComputation independent_functionElimination dependent_pairFormation_alt lambdaEquality_alt int_eqEquality independent_pairFormation universeIsType lambdaFormation_alt because_Cache productIsType instantiate cumulativity intEquality equalityTransitivity equalitySymmetry inhabitedIsType equalityElimination productElimination equalityIsType2 baseApply closedConclusion baseClosed applyEquality promote_hyp int_eqReduceTrueSq equalityIsType1 isect_memberFormation_alt axiomSqEquality sqequalSqle divergentSqle callbyvalueLess lessCases imageMemberEquality imageElimination sqleReflexivity lessExceptionCases axiomSqleEquality exceptionSqequal exceptionLess

Latex:
\mforall{}[F,G1,G2,L1,L2,a:Top].  \mforall{}[m1,m2:\mBbbN{}].
    (cbva\_seq(\mlambda{}n.if  (n)  <  (m1)
                                    then  L1  a  n
                                    else  mk\_lambdas\_fun(\mlambda{}g.if  n  -  m1=m2
                                                                                  then  mk\_lambdas\_fun(\mlambda{}g1.((G1  g)  +  (G2  g1));m2)
                                                                                  else  (L2  a  (n  -  m1));m1);  F;  (m1  +  m2)  +  1) 
    \msim{}  cbva\_seq(\mlambda{}n.if  (n)  <  (m1)
                                      then  L1  a  n
                                      else  if  (n)  <  (m1  +  m2)
                                                      then  mk\_lambdas(L2  a  (n  -  m1);m1)
                                                      else  mk\_lambdas\_fun(\mlambda{}g1.mk\_lambdas\_fun(\mlambda{}g2.((G1  g1)  +  (G2  g2));m2);m1);
                          F;  (m1  +  m2)  +  1))



Date html generated: 2019_10_15-AM-10_58_50
Last ObjectModification: 2018_10_11-PM-09_50_58

Theory : untyped!computation


Home Index