]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/reduction/cpr.ma
- new component "s_transition" for the restored fqu and fquq
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / reduction / cpr.ma
index 61d8fd9317d55a6493b29ccc37b8f8e78d2ca963..d5374513aa6a6d8aff14b368ba9fe7cd1e9c17a2 100644 (file)
@@ -81,7 +81,7 @@ lemma cpr_delift: ∀G,K,V,T1,L,l. ⬇[l] L ≡ (K.ⓓV) →
 #G #K #V #T1 elim T1 -T1
 [ * /2 width=4 by cpr_atom, lift_sort, lift_gref, ex2_2_intro/
   #i #L #l #HLK elim (lt_or_eq_or_gt i l)
-  #Hil [1,3: /3 width=4 by cpr_atom, lift_lref_ge_minus, lift_lref_lt, ex2_2_intro/ ]
+  #Hil [1,3: /4 width=4 by lift_lref_ge_minus, lift_lref_lt, ylt_inj, yle_inj, ex2_2_intro/ ]
   destruct
   elim (lift_total V 0 (i+1)) #W #HVW
   elim (lift_split … HVW i i) /3 width=6 by cpr_delta, ex2_2_intro/
@@ -97,8 +97,7 @@ fact lstas_cpr_aux: ∀h,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 •*[h, d] T2 →
                     d = 0 → ⦃G, L⦄ ⊢ T1 ➡ T2.
 #h #G #L #T1 #T2 #d #H elim H -G -L -T1 -T2 -d
 /3 width=1 by cpr_eps, cpr_flat, cpr_bind/
-[ #G #L #d #k #H0 destruct normalize //
-| #G #L #K #V1 #V2 #W2 #i #d #HLK #_ #HVW2 #IHV12 #H destruct
+[ #G #L #K #V1 #V2 #W2 #i #d #HLK #_ #HVW2 #IHV12 #H destruct
   /3 width=6 by cpr_delta/
 | #G #L #K #V1 #V2 #W2 #i #d #_ #_ #_ #_ <plus_n_Sm #H destruct
 ]
@@ -132,8 +131,8 @@ lemma cpr_inv_atom1: ∀I,G,L,T2. ⦃G, L⦄ ⊢ ⓪{I} ➡ T2 →
 /2 width=3 by cpr_inv_atom1_aux/ qed-.
 
 (* Basic_1: includes: pr0_gen_sort pr2_gen_sort *)
-lemma cpr_inv_sort1: ∀G,L,T2,k. ⦃G, L⦄ ⊢ ⋆k ➡ T2 → T2 = ⋆k.
-#G #L #T2 #k #H
+lemma cpr_inv_sort1: ∀G,L,T2,s. ⦃G, L⦄ ⊢ ⋆s ➡ T2 → T2 = ⋆s.
+#G #L #T2 #s #H
 elim (cpr_inv_atom1 … H) -H //
 * #K #V #V2 #i #_ #_ #_ #H destruct
 qed-.
@@ -264,7 +263,7 @@ lemma cpr_inv_appl1_simple: ∀G,L,V1,T1,U. ⦃G, L⦄ ⊢ ⓐV1. T1 ➡ U → 
                                      U = ⓐV2. T2.
 #G #L #V1 #T1 #U #H #HT1
 elim (cpr_inv_appl1 … H) -H *
-[ /2 width=5/
+[ /2 width=5 by ex3_2_intro/
 | #a #V2 #W1 #W2 #U1 #U2 #_ #_ #_ #H #_ destruct
   elim (simple_inv_bind … HT1)
 | #a #V #V2 #W1 #W2 #U1 #U2 #_ #_ #_ #_ #H #_ destruct