]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/computation/lprs_cprs.ma
commit completed: now we support two versions of slicing for local
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / computation / lprs_cprs.ma
index a1c4086603883fa52155bd5291d4a0656765cb51..70345be5c9dbbecc226318945fadd70f11f37ba7 100644 (file)
@@ -49,13 +49,13 @@ lemma lprs_cprs_conf_dx: ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡* T1 →
                          ∃∃T. ⦃G, L1⦄ ⊢ T1 ➡* T & ⦃G, L1⦄ ⊢ T0 ➡* T.
 #G #L0 #T0 #T1 #HT01 #L1 #H elim H -L1
 [ #L1 #HL01
-  elim (cprs_lpr_conf_dx … HT01 … HL01) -L0 /2 width=3/
+  elim (cprs_lpr_conf_dx … HT01 … HL01) -L0 /2 width=3 by ex2_intro/
 | #L #L1 #_ #HL1 * #T #HT1 #HT0 -L0
   elim (cprs_lpr_conf_dx … HT1 … HL1) -HT1 #T2 #HT2 #HT12
   elim (cprs_lpr_conf_dx … HT0 … HL1) -L #T3 #HT3 #HT03
   elim (cprs_conf … HT2 … HT3) -T #T #HT2 #HT3
   lapply (cprs_trans … HT03 … HT3) -T3
-  lapply (cprs_trans … HT12 … HT2) -T2 /2 width=3/
+  lapply (cprs_trans … HT12 … HT2) -T2 /2 width=3 by ex2_intro/
 ]
 qed-.
 
@@ -69,7 +69,7 @@ lemma lprs_cprs_conf_sn: ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡* T1 →
                          ∃∃T. ⦃G, L0⦄ ⊢ T1 ➡* T & ⦃G, L1⦄ ⊢ T0 ➡* T.
 #G #L0 #T0 #T1 #HT01 #L1 #HL01
 elim (lprs_cprs_conf_dx … HT01 … HL01) -HT01 #T #HT1
-lapply (lprs_cprs_trans … HT1 … HL01) -HT1 /2 width=3/
+lapply (lprs_cprs_trans … HT1 … HL01) -HT1 /2 width=3 by ex2_intro/
 qed-.
 
 lemma lprs_cpr_conf_sn: ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡ T1 →
@@ -81,7 +81,7 @@ lemma cprs_bind2: ∀G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡* V2 →
                   ∀I,T1,T2. ⦃G, L.ⓑ{I}V2⦄ ⊢ T1 ➡* T2 →
                   ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡* ⓑ{a,I}V2.T2.
 #G #L #V1 #V2 #HV12 #I #T1 #T2 #HT12
-lapply (lprs_cprs_trans … HT12 (L.ⓑ{I}V1) ?) /2 width=1/
+lapply (lprs_cprs_trans … HT12 (L.ⓑ{I}V1) ?) /2 width=1 by lprs_pair, cprs_bind/
 qed.
 
 (* Inversion lemmas on context-sensitive parallel computation for terms *****)
@@ -90,18 +90,18 @@ qed.
 lemma cprs_inv_abst1: ∀a,G,L,W1,T1,U2. ⦃G, L⦄ ⊢ ⓛ{a}W1.T1 ➡* U2 →
                       ∃∃W2,T2. ⦃G, L⦄ ⊢ W1 ➡* W2 & ⦃G, L.ⓛW1⦄ ⊢ T1 ➡* T2 &
                                U2 = ⓛ{a}W2.T2.
-#a #G #L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /2 width=5/
+#a #G #L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /2 width=5 by ex3_2_intro/
 #U0 #U2 #_ #HU02 * #V0 #T0 #HV10 #HT10 #H destruct
 elim (cpr_inv_abst1 … HU02) -HU02 #V2 #T2 #HV02 #HT02 #H destruct
-lapply (lprs_cpr_trans … HT02 (L.ⓛV1) ?) /2 width=1/ -HT02 #HT02
+lapply (lprs_cpr_trans … HT02 (L.ⓛV1) ?) /2 width=1 by lprs_pair/ -HT02 #HT02
 lapply (cprs_strap1 … HV10 … HV02) -V0
-lapply (cprs_trans … HT10 … HT02) -T0 /2 width=5/
+lapply (cprs_trans … HT10 … HT02) -T0 /2 width=5 by ex3_2_intro/
 qed-.
 
 lemma cprs_inv_abst: ∀a,G,L,W1,W2,T1,T2. ⦃G, L⦄ ⊢ ⓛ{a}W1.T1 ➡* ⓛ{a}W2.T2 →
                      ⦃G, L⦄ ⊢ W1 ➡* W2 ∧ ⦃G, L.ⓛW1⦄ ⊢ T1 ➡* T2.
 #a #G #L #W1 #W2 #T1 #T2 #H
-elim (cprs_inv_abst1 … H) -H #W #T #HW1 #HT1 #H destruct /2 width=1/
+elim (cprs_inv_abst1 … H) -H #W #T #HW1 #HT1 #H destruct /2 width=1 by conj/
 qed-.
 
 (* Basic_1: was pr3_gen_abbr *)
@@ -110,21 +110,22 @@ lemma cprs_inv_abbr1: ∀a,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓓ{a}V1.T1 ➡* U2 →
                                U2 = ⓓ{a}V2.T2
                       ) ∨
                       ∃∃T2. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡* T2 & ⇧[0, 1] U2 ≡ T2 & a = true.
-#a #G #L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /3 width=5/
+#a #G #L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /3 width=5 by ex3_2_intro, or_introl/
 #U0 #U2 #_ #HU02 * *
 [ #V0 #T0 #HV10 #HT10 #H destruct
   elim (cpr_inv_abbr1 … HU02) -HU02 *
   [ #V2 #T2 #HV02 #HT02 #H destruct
-    lapply (lprs_cpr_trans … HT02 (L.ⓓV1) ?) /2 width=1/ -HT02 #HT02
+    lapply (lprs_cpr_trans … HT02 (L.ⓓV1) ?) /2 width=1 by lprs_pair/ -HT02 #HT02
     lapply (cprs_strap1 … HV10 … HV02) -V0
-    lapply (cprs_trans … HT10 … HT02) -T0 /3 width=5/
+    lapply (cprs_trans … HT10 … HT02) -T0 /3 width=5 by ex3_2_intro, or_introl/
   | #T2 #HT02 #HUT2
-    lapply (lprs_cpr_trans … HT02 (L.ⓓV1) ?) -HT02 /2 width=1/ -V0 #HT02
-    lapply (cprs_trans … HT10 … HT02) -T0 /3 width=3/
+    lapply (lprs_cpr_trans … HT02 (L.ⓓV1) ?) -HT02 /2 width=1 by lprs_pair/ -V0 #HT02
+    lapply (cprs_trans … HT10 … HT02) -T0 /3 width=3 by ex3_intro, or_intror/
   ]
 | #U1 #HTU1 #HU01
   elim (lift_total U2 0 1) #U #HU2
-  lapply (cpr_lift … HU02 (L.ⓓV1) … HU01 … HU2) -U0 /2 width=1/ /4 width=3/
+  lapply (cpr_lift … HU02 (L.ⓓV1) … HU01 … HU2) -U0
+  /4 width=3 by cprs_strap1, ldrop_drop, ex3_intro, or_intror/
 ]
 qed-.