]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/computation/csn_lpx.ma
- probe: critical bug fixed (all objects were deleted due to wrong test)
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / computation / csn_lpx.ma
index 7fcc3fa0e7e6f55d39a7580cf8cd662c59ae96c8..3f97c882add74d5617d4c41f3047c156a6d27789 100644 (file)
@@ -58,43 +58,26 @@ elim (cpx_inv_abbr1 … H1) -H1 *
 ]
 qed.
 
-fact csn_appl_beta_aux: ∀h,g,a,L,U1. ⦃h, L⦄ ⊢ ⬊*[g] U1 → ∀V. ⦃h, L⦄ ⊢ ⬊*[g] V →
-                        ∀W,T1. U1 = ⓛ{a}W.T1 → (
-                           ∀T2. ⦃h, L.ⓛW⦄ ⊢ T1 ➡*[g] T2 → ⦃h, L⦄ ⊢ ⬊*[g] ⓓ{a}V.T2
-                        ) → ⦃h, L⦄ ⊢ ⬊*[g] ⓐV.ⓛ{a}W.T1.
-#h #g #a #L #X1 #H @(csn_ind … H) -X1
-#X1 #HT1 #IHT1 #X2 #H @(csn_ind … H) -X2
-#V #HV #IHV #W #T1 #H1 #IHT2 destruct
+fact csn_appl_beta_aux: ∀h,g,a,L,U1. ⦃h, L⦄ ⊢ ⬊*[g] U1 →
+                        ∀V,W,T1. U1 = ⓓ{a}ⓝW.V.T1 → ⦃h, L⦄ ⊢ ⬊*[g] ⓐV.ⓛ{a}W.T1.
+#h #g #a #L #X #H @(csn_ind … H) -X
+#X #HT1 #IHT1 #V #W #T1 #H1 destruct
 @csn_intro #X #H1 #H2
 elim (cpx_inv_appl1 … H1) -H1 *
-[ #V0 #Y #HLV0 #H #H0 destruct
+[ -HT1 #V0 #Y #HLV0 #H #H0 destruct
   elim (cpx_inv_abst1 … H) -H #W0 #T0 #HLW0 #HLT0 #H destruct
-  elim (eq_false_inv_tpair_dx … H2) -H2
-  [ lapply (csn_cpx_trans … HV … HLV0) -HV #HV0 #HWT0
-    @IHT1 -IHT1 [4,5: // |1: skip |2,3: /2 width=1/ ] -HWT0 -HV0 #T2 #HT02
-    lapply (lpx_cpxs_trans … HT02 (L.ⓛW) ?) [ /2 width=1/ ] -W0 #HT02
-    lapply (cpxs_strap2 … HLT0 … HT02) -T0 #HT12
-    lapply (IHT2 … HT12) -T1 #HT2
-    @(csn_cpx_trans … HT2) -HT2 /2 width=1/
-  | -HV -HT1 -IHT1 -HLW0 -HLT0 * #H #HV0 destruct
-    @IHV -IHV [1,3: // |2: /2 width=1/ ] -HV0 #T2 #HT02
-    lapply (IHT2 … HT02) -IHT2 -HT02 #HT2
-    @(csn_cpx_trans … HT2) -HT2 /2 width=1/
-  ]
-| -HT1 -IHT1 -HV -IHV -H2 #b #V0 #W0 #T0 #T3 #HLV0 #HLT01 #H1 #H2 destruct
-  lapply (IHT2 T3 ?) [ /2 width=1/ ] -IHT2 -HLT01 #HT3
-  @(csn_cpx_trans … HT3) -HT3 /2 width=1/
-| -HT1 -IHT1 -HV -IHV -IHT2 -H2 #b #V0 #V1 #W0 #W1 #T0 #T3 #_ #_ #_ #_ #H destruct
+  @IHT1 -IHT1 [4: // | skip |3: #H destruct /2 width=1/ ] -H2
+  lapply (lsubr_cpx_trans … HLT0 (L.ⓓⓝW.V) ?) -HLT0 [ /2 width=1/ ] /3 width=1/
+| -IHT1 -H2 #b #V0 #W0 #W2 #T0 #T2 #HLV0 #HLW02 #HLT02 #H1 #H3 destruct
+  lapply (lsubr_cpx_trans … HLT02 (L.ⓓⓝW0.V) ?) -HLT02 [ /2 width=1/ ] #HT02
+  @(csn_cpx_trans … HT1) -HT1 /3 width=1/
+| -HT1 -IHT1 -H2 #b #V0 #V1 #W0 #W1 #T0 #T3 #_ #_ #_ #_ #H destruct
 ]
 qed-.
 
 (* Basic_1: was just: sn3_beta *)
-lemma csn_appl_beta: ∀h,g,a,L,W,T1. ⦃h, L⦄ ⊢ ⬊*[g] ⓛ{a}W.T1 → ∀V. (
-                        ∀T2. ⦃h, L.ⓛW⦄ ⊢ T1 ➡*[g] T2 → ⦃h, L⦄ ⊢ ⬊*[g] ⓓ{a}V.T2
-                     ) → ⦃h, L⦄ ⊢ ⬊*[g] ⓐV.ⓛ{a}W.T1.
-#h #g #a #L #W #T1 #HWT1 #V #IHT2 lapply (IHT2 T1 ?) //
-#HVT1 lapply (csn_fwd_pair_sn … HVT1) -HVT1
-/3 width=3 by csn_appl_beta_aux/ qed.
+lemma csn_appl_beta: ∀h,g,a,L,V,W,T. ⦃h, L⦄ ⊢ ⬊*[g] ⓓ{a}ⓝW.V.T → ⦃h, L⦄ ⊢ ⬊*[g] ⓐV.ⓛ{a}W.T.
+/2 width=3 by csn_appl_beta_aux/ qed.
 
 fact csn_appl_theta_aux: ∀h,g,a,L,U. ⦃h, L⦄ ⊢ ⬊*[g] U → ∀V1,V2. ⇧[0, 1] V1 ≡ V2 →
                          ∀V,T. U = ⓓ{a}V.ⓐV2.T → ⦃h, L⦄ ⊢ ⬊*[g] ⓐV1.ⓓ{a}V.T.
@@ -124,7 +107,7 @@ elim (cpx_inv_appl1 … HL) -HL *
     lapply (csn_inv_lift … L … 1 HVT0 ? ? ?) -HVT0 [ /2 width=4/ |2,3: skip | /2 width=1/ ] -V2 -T0 #HVY
     @(csn_cpx_trans … HVY) /2 width=1/
   ]
-| -HV -HV12 -HVT -IHVT -H #b #V0 #W0 #T0 #T1 #_ #_ #H destruct
+| -HV -HV12 -HVT -IHVT -H #b #V0 #W0 #W1 #T0 #T1 #_ #_ #_ #H destruct
 | -IHVT -H #b #V0 #V3 #W0 #W1 #T0 #T1 #HLV10 #HV03 #HLW01 #HLT01 #H1 #H2 destruct
   lapply (cpx_lift … HLV10 (L. ⓓW0) … HV12 … HV03) -HLV10 -HV12 -HV03 /2 width=1/ #HLV23
   @csn_abbr /2 width=3 by csn_cpx_trans/ -HV
@@ -143,7 +126,7 @@ lemma csn_appl_simple_tstc: ∀h,g,L,V. ⦃h, L⦄ ⊢ ⬊*[g] V → ∀T1. ⦃h
                             𝐒⦃T1⦄ → ⦃h, L⦄ ⊢ ⬊*[g] ⓐV.T1.
 #h #g #L #V #H @(csn_ind … H) -V #V #_ #IHV #T1 #H @(csn_ind … H) -T1 #T1 #H1T1 #IHT1 #H2T1 #H3T1
 @csn_intro #X #HL #H
-elim (cpx_inv_appl1_simple … HL ?) -HL //
+elim (cpx_inv_appl1_simple … HL) -HL //
 #V0 #T0 #HLV0 #HLT10 #H0 destruct
 elim (eq_false_inv_tpair_sn … H) -H
 [ -IHT1 #HV0