]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/reduction/cnx.ma
- commit of the "reduction" component with two additions ...
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / reduction / cnx.ma
index ba3c976ac8cf7c5ee9a69200b15071c5ed9305a0..9954e6edbc2795eff61fc78bd5b29fba62de7654 100644 (file)
@@ -30,41 +30,42 @@ interpretation
 lemma cnx_inv_sort: ∀h,g,G,L,k. ⦃G, L⦄ ⊢ 𝐍[h, g]⦃⋆k⦄ → deg h g k 0.
 #h #g #G #L #k #H elim (deg_total h g k)
 #l @(nat_ind_plus … l) -l // #l #_ #Hkl
-lapply (H (⋆(next h k)) ?) -H /2 width=2/ -L -l #H destruct -H -e0 (**) (* destruct does not remove some premises *)
+lapply (H (⋆(next h k)) ?) -H /2 width=2 by cpx_sort/ -L -l #H destruct -H -e0 (**) (* destruct does not remove some premises *)
 lapply (next_lt h k) >e1 -e1 #H elim (lt_refl_false … H)
 qed-.
 
-lemma cnx_inv_delta: ∀h,g,I,G,L,K,V,i. ⇩[0, i] L ≡ K.ⓑ{I}V → ⦃G, L⦄ ⊢ 𝐍[h, g]⦃#i⦄ → ⊥.
+lemma cnx_inv_delta: ∀h,g,I,G,L,K,V,i. ⇩[i] L ≡ K.ⓑ{I}V → ⦃G, L⦄ ⊢ 𝐍[h, g]⦃#i⦄ → ⊥.
 #h #g #I #G #L #K #V #i #HLK #H
 elim (lift_total V 0 (i+1)) #W #HVW
-lapply (H W ?) -H [ /3 width=7/ ] -HLK #H destruct
+lapply (H W ?) -H [ /3 width=7 by cpx_delta/ ] -HLK #H destruct
 elim (lift_inv_lref2_be … HVW) -HVW //
 qed-.
 
 lemma cnx_inv_abst: ∀h,g,a,G,L,V,T. ⦃G, L⦄ ⊢ 𝐍[h, g]⦃ⓛ{a}V.T⦄ →
                     ⦃G, L⦄ ⊢ 𝐍[h, g]⦃V⦄ ∧ ⦃G, L.ⓛV⦄ ⊢ 𝐍[h, g]⦃T⦄.
 #h #g #a #G #L #V1 #T1 #HVT1 @conj
-[ #V2 #HV2 lapply (HVT1 (ⓛ{a}V2.T1) ?) -HVT1 /2 width=2/ -HV2 #H destruct //
-| #T2 #HT2 lapply (HVT1 (ⓛ{a}V1.T2) ?) -HVT1 /2 width=2/ -HT2 #H destruct //
+[ #V2 #HV2 lapply (HVT1 (ⓛ{a}V2.T1) ?) -HVT1 /2 width=2 by cpx_pair_sn/ -HV2 #H destruct //
+| #T2 #HT2 lapply (HVT1 (ⓛ{a}V1.T2) ?) -HVT1 /2 width=2 by cpx_bind/ -HT2 #H destruct //
 ]
 qed-.
 
 lemma cnx_inv_abbr: ∀h,g,G,L,V,T. ⦃G, L⦄ ⊢ 𝐍[h, g]⦃-ⓓV.T⦄ →
                     ⦃G, L⦄ ⊢ 𝐍[h, g]⦃V⦄ ∧ ⦃G, L.ⓓV⦄ ⊢ 𝐍[h, g]⦃T⦄.
 #h #g #G #L #V1 #T1 #HVT1 @conj
-[ #V2 #HV2 lapply (HVT1 (-ⓓV2.T1) ?) -HVT1 /2 width=2/ -HV2 #H destruct //
-| #T2 #HT2 lapply (HVT1 (-ⓓV1.T2) ?) -HVT1 /2 width=2/ -HT2 #H destruct //
+[ #V2 #HV2 lapply (HVT1 (-ⓓV2.T1) ?) -HVT1 /2 width=2 by cpx_pair_sn/ -HV2 #H destruct //
+| #T2 #HT2 lapply (HVT1 (-ⓓV1.T2) ?) -HVT1 /2 width=2 by cpx_bind/ -HT2 #H destruct //
 ]
 qed-.
 
 lemma cnx_inv_zeta: ∀h,g,G,L,V,T. ⦃G, L⦄ ⊢ 𝐍[h, g]⦃+ⓓV.T⦄ → ⊥.
 #h #g #G #L #V #T #H elim (is_lift_dec T 0 1)
 [ * #U #HTU
-  lapply (H U ?) -H /2 width=3/ #H destruct
+  lapply (H U ?) -H /2 width=3 by cpx_zeta/ #H destruct
   elim (lift_inv_pair_xy_y … HTU)
 | #HT
   elim (cpr_delift G(⋆) V T (⋆.ⓓV) 0) // #T2 #T1 #HT2 #HT12
-  lapply (H (+ⓓV.T2) ?) -H /5 width=1/ -HT2 #H destruct /3 width=2/
+  lapply (H (+ⓓV.T2) ?) -H /5 width=1 by cpr_cpx, tpr_cpr, cpr_bind/ -HT2
+  #H destruct /3 width=2 by ex_intro/
 ]
 qed-.
 
@@ -77,19 +78,20 @@ lemma cnx_inv_appl: ∀h,g,G,L,V,T. ⦃G, L⦄ ⊢ 𝐍[h, g]⦃ⓐV.T⦄ →
   [ elim (lift_total V1 0 1) #V2 #HV12
     lapply (H (ⓓ{a}W1.ⓐV2.U1) ?) -H /3 width=3/ -HV12 #H destruct
   | lapply (H (ⓓ{a}ⓝW1.V1.U1) ?) -H /3 width=1/ #H destruct
+  ]
 ]
 qed-.
 
 lemma cnx_inv_tau: ∀h,g,G,L,V,T. ⦃G, L⦄ ⊢ 𝐍[h, g]⦃ⓝV.T⦄ → ⊥.
-#h #g #G #L #V #T #H lapply (H T ?) -H /2 width=1/ #H
-@discr_tpair_xy_y //
+#h #g #G #L #V #T #H lapply (H T ?) -H
+/2 width=4 by cpx_tau, discr_tpair_xy_y/
 qed-.
 
 (* Basic forward lemmas *****************************************************)
 
 lemma cnx_fwd_cnr: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ 𝐍[h, g]⦃T⦄ → ⦃G, L⦄ ⊢ 𝐍⦃T⦄.
 #h #g #G #L #T #H #U #HTU
-@H /2 width=1/ (**) (* auto fails because a δ-expansion gets in the way *)
+@H /2 width=1 by cpr_cpx/ (**) (* auto fails because a δ-expansion gets in the way *)
 qed-.
 
 (* Basic properties *********************************************************)
@@ -101,7 +103,7 @@ qed.
 
 lemma cnx_sort_iter: ∀h,g,G,L,k,l. deg h g k l → ⦃G, L⦄ ⊢ 𝐍[h, g]⦃⋆((next h)^l k)⦄.
 #h #g #G #L #k #l #Hkl
-lapply (deg_iter … l Hkl) -Hkl <minus_n_n /2 width=1/
+lapply (deg_iter … l Hkl) -Hkl <minus_n_n /2 width=1 by cnx_sort/
 qed.
 
 lemma cnx_abst: ∀h,g,a,G,L,W,T. ⦃G, L⦄ ⊢ 𝐍[h, g]⦃W⦄ → ⦃G, L.ⓛW⦄ ⊢ 𝐍[h, g]⦃T⦄ →