]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_simple.ma
update in ground_2, static_2, basic_2, apps_2, alpha_1
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_transition / cpx_simple.ma
index fbc9eb6442cfaf2259db5809bc5cc2f0abe469ea..d0b3e453e1ab52d7137604408a3e06e4e1bc8331 100644 (file)
 include "basic_2/rt_transition/cpg_simple.ma".
 include "basic_2/rt_transition/cpx.ma".
 
-(* UNCOUNTED CONTEXT-SENSITIVE PARALLEL RT-TRANSITION FOR TERMS *************)
+(* UNBOUND CONTEXT-SENSITIVE PARALLEL RT-TRANSITION FOR TERMS ***************)
 
-lemma cpx_inv_appl1_simple: ∀h,G,L,V1,T1,U. ⦃G, L⦄ ⊢ ⓐV1.T1 ⬈[h] U → 𝐒⦃T1⦄ →
-                            ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ⬈[h] V2 & ⦃G, L⦄ ⊢ T1 ⬈[h] T2 &
+(* Inversion lemmas with simple terms ***************************************)
+
+lemma cpx_inv_appl1_simple: ∀h,G,L,V1,T1,U. ❪G,L❫ ⊢ ⓐV1.T1 ⬈[h] U → 𝐒❪T1❫ →
+                            ∃∃V2,T2. ❪G,L❫ ⊢ V1 ⬈[h] V2 & ❪G,L❫ ⊢ T1 ⬈[h] T2 &
                                      U = ⓐV2.T2.
 #h #G #L #V1 #T1 #U * #c #H #HT1 elim (cpg_inv_appl1_simple … H) -H
 /3 width=5 by ex3_2_intro, ex_intro/