Nuprl Lemma : es-local-pred-property2
∀[Info:Type]
∀es:EO+(Info). ∀e:E. ∀P:{e':E| (e' <loc e)} ⟶ 𝔹.
((↑can-apply(last(P);e)
⇐⇒ ∃a:E. ((a <loc e) ∧ (↑(P a))))
∧ (do-apply(last(P);e) <loc e)
∧ (↑(P do-apply(last(P);e)))
∧ (∀e'':E. ((e'' <loc e)
⇒ (do-apply(last(P);e) <loc e'')
⇒ (¬↑(P e''))))
supposing ↑can-apply(last(P);e))
Proof
Definitions occuring in Statement :
es-local-pred: last(P)
,
event-ordering+: EO+(Info)
,
es-locl: (e <loc e')
,
es-E: E
,
assert: ↑b
,
bool: 𝔹
,
uimplies: b supposing a
,
uall: ∀[x:A]. B[x]
,
all: ∀x:A. B[x]
,
exists: ∃x:A. B[x]
,
iff: P
⇐⇒ Q
,
not: ¬A
,
implies: P
⇒ Q
,
and: P ∧ Q
,
set: {x:A| B[x]}
,
apply: f a
,
function: x:A ⟶ B[x]
,
universe: Type
,
do-apply: do-apply(f;x)
,
can-apply: can-apply(f;x)
Definitions unfolded in proof :
uall: ∀[x:A]. B[x]
,
member: t ∈ T
,
all: ∀x:A. B[x]
,
do-apply: do-apply(f;x)
,
can-apply: can-apply(f;x)
,
subtype_rel: A ⊆r B
,
prop: ℙ
,
so_lambda: λ2x.t[x]
,
and: P ∧ Q
,
implies: P
⇒ Q
,
so_apply: x[s]
,
or: P ∨ Q
,
isl: isl(x)
,
outl: outl(x)
,
assert: ↑b
,
ifthenelse: if b then t else f fi
,
btrue: tt
,
cand: A c∧ B
,
iff: P
⇐⇒ Q
,
rev_implies: P
⇐ Q
,
true: True
,
uimplies: b supposing a
,
sq_exists: ∃x:{A| B[x]}
,
sq_type: SQType(T)
,
guard: {T}
,
not: ¬A
,
false: False
,
es-locl: (e <loc e')
,
es-causl: (e < e')
,
squash: ↓T
,
bfalse: ff
,
exists: ∃x:A. B[x]
,
sq_stable: SqStable(P)
Latex:
\mforall{}[Info:Type]
\mforall{}es:EO+(Info). \mforall{}e:E. \mforall{}P:\{e':E| (e' <loc e)\} {}\mrightarrow{} \mBbbB{}.
((\muparrow{}can-apply(last(P);e) \mLeftarrow{}{}\mRightarrow{} \mexists{}a:E. ((a <loc e) \mwedge{} (\muparrow{}(P a))))
\mwedge{} (do-apply(last(P);e) <loc e)
\mwedge{} (\muparrow{}(P do-apply(last(P);e)))
\mwedge{} (\mforall{}e'':E. ((e'' <loc e) {}\mRightarrow{} (do-apply(last(P);e) <loc e'') {}\mRightarrow{} (\mneg{}\muparrow{}(P e''))))
supposing \muparrow{}can-apply(last(P);e))
Date html generated:
2016_05_16-PM-11_27_23
Last ObjectModification:
2016_01_17-PM-07_10_57
Theory : event-ordering
Home
Index