Nuprl Lemma : do-apply-p-first

[A,B:Type]. ∀[L:(A ⟶ (B Top)) List]. ∀[x:A].
  do-apply(p-first(L);x) do-apply(hd(filter(λf.can-apply(f;x);L));x) ∈ supposing ↑can-apply(p-first(L);x)


Proof




Definitions occuring in Statement :  p-first: p-first(L) do-apply: do-apply(f;x) can-apply: can-apply(f;x) hd: hd(l) filter: filter(P;l) list: List assert: b uimplies: supposing a uall: [x:A]. B[x] top: Top lambda: λx.A[x] function: x:A ⟶ B[x] union: left right universe: Type equal: t ∈ T
Definitions unfolded in proof :  uall: [x:A]. B[x] member: t ∈ T all: x:A. B[x] nat: implies:  Q false: False ge: i ≥  uimplies: supposing a satisfiable_int_formula: satisfiable_int_formula(fmla) exists: x:A. B[x] not: ¬A top: Top and: P ∧ Q prop: subtype_rel: A ⊆B guard: {T} or: P ∨ Q cons: [a b] colength: colength(L) so_lambda: λ2y.t[x; y] so_apply: x[s1;s2] decidable: Dec(P) so_lambda: λ2x.t[x] so_apply: x[s] nil: [] it: sq_type: SQType(T) less_than: a < b squash: T less_than': less_than'(a;b) assert: b ifthenelse: if then else fi  bfalse: ff append: as bs so_lambda: so_lambda(x,y,z.t[x; y; z]) so_apply: x[s1;s2;s3] true: True uiff: uiff(P;Q) iff: ⇐⇒ Q bool: 𝔹 unit: Unit btrue: tt do-apply: do-apply(f;x) p-first: p-first(L) can-apply: can-apply(f;x) isl: isl(x) outl: outl(x) rev_implies:  Q rev_uimplies: rev_uimplies(P;Q) p-conditional: [f?g]
Lemmas referenced :  nat_properties satisfiable-full-omega-tt intformand_wf intformle_wf itermConstant_wf itermVar_wf intformless_wf int_formula_prop_and_lemma int_formula_prop_le_lemma int_term_value_constant_lemma int_term_value_var_lemma int_formula_prop_less_lemma int_formula_prop_wf ge_wf less_than_wf assert_wf can-apply_wf p-first_wf top_wf less_than_transitivity1 less_than_irreflexivity equal-wf-T-base nat_wf colength_wf_list list-cases product_subtype_list spread_cons_lemma intformeq_wf itermAdd_wf int_formula_prop_eq_lemma int_term_value_add_lemma decidable__le intformnot_wf int_formula_prop_not_lemma le_wf equal_wf subtype_rel_list subtype_rel_dep_function subtype_rel_union subtract_wf itermSubtract_wf int_term_value_subtract_lemma subtype_base_sq set_subtype_base int_subtype_base decidable__equal_int list_wf p_first_nil_lemma filter_nil_lemma false_wf list_ind_cons_lemma list_ind_nil_lemma assert_functionality_wrt_uiff cons_wf squash_wf true_wf p-first-append nil_wf p-conditional_wf p-conditional-to-p-first p-conditional-domain decidable__assert append_wf filter_cons_lemma bool_wf bnot_wf not_wf eqtt_to_assert uiff_transitivity eqff_to_assert assert_of_bnot reduce_hd_cons_lemma list_accum_cons_lemma list_induction all_wf list_accum_wf list_accum_nil_lemma p-first-singleton iff_weakening_equal do-apply_wf outl_wf
Rules used in proof :  sqequalSubstitution sqequalTransitivity computationStep sqequalReflexivity isect_memberFormation introduction cut thin lambdaFormation extract_by_obid sqequalHypSubstitution isectElimination hypothesisEquality hypothesis setElimination rename intWeakElimination natural_numberEquality independent_isectElimination dependent_pairFormation lambdaEquality int_eqEquality intEquality dependent_functionElimination isect_memberEquality voidElimination voidEquality sqequalRule independent_pairFormation computeAll independent_functionElimination axiomEquality because_Cache applyEquality equalityTransitivity equalitySymmetry functionEquality cumulativity unionEquality unionElimination promote_hyp hypothesis_subsumption productElimination applyLambdaEquality dependent_set_memberEquality addEquality baseClosed instantiate imageElimination universeEquality functionExtensionality imageMemberEquality equalityElimination inlEquality hyp_replacement inrFormation

Latex:
\mforall{}[A,B:Type].  \mforall{}[L:(A  {}\mrightarrow{}  (B  +  Top))  List].  \mforall{}[x:A].
    do-apply(p-first(L);x)  =  do-apply(hd(filter(\mlambda{}f.can-apply(f;x);L));x) 
    supposing  \muparrow{}can-apply(p-first(L);x)



Date html generated: 2018_05_21-PM-06_44_41
Last ObjectModification: 2017_07_26-PM-04_55_12

Theory : general


Home Index