]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/Basic_2/reducibility/cpr.ma
commit by user andrea
[helm.git] / matita / matita / contribs / lambda_delta / Basic_2 / reducibility / cpr.ma
index bcc50fa6e8f6c68e464b831a7f39a5f51250e2c2..0ad09e4b4c22a1bd0dec07b9ed2645d2a143a073 100644 (file)
@@ -28,6 +28,9 @@ interpretation
 
 (* Basic properties *********************************************************)
 
+lemma cpr_intro: ∀L,T1,T,T2,d,e. T1 ➡ T → L ⊢ T [d, e] ▶* T2 → L ⊢ T1 ➡ T2.
+/4 width=3/ qed-.
+
 (* Basic_1: was by definition: pr2_free *)
 lemma cpr_pr: ∀T1,T2. T1 ➡ T2 → ∀L. L ⊢ T1 ➡ T2.
 /2 width=3/ qed.
@@ -41,12 +44,12 @@ lemma cpr_refl: ∀L,T. L ⊢ T ➡ T.
 (* Note: new property *)
 (* Basic_1: was only: pr2_thin_dx *) 
 lemma cpr_flat: ∀I,L,V1,V2,T1,T2.
-                L ⊢ V1 ➡ V2 → L ⊢ T1 ➡ T2 → L ⊢ 𝕗{I} V1. T1 ➡ 𝕗{I} V2. T2.
+                L ⊢ V1 ➡ V2 → L ⊢ T1 ➡ T2 → L ⊢ ⓕ{I} V1. T1 ➡ ⓕ{I} V2. T2.
 #I #L #V1 #V2 #T1 #T2 * #V #HV1 #HV2 * /3 width=5/
 qed.
 
 lemma cpr_cast: ∀L,V,T1,T2.
-                L ⊢ T1 ➡ T2 → L ⊢ 𝕔{Cast} V. T1 ➡ T2.
+                L ⊢ T1 ➡ T2 → L ⊢ V. T1 ➡ T2.
 #L #V #T1 #T2 * /3 width=3/
 qed.
 
@@ -73,10 +76,30 @@ lemma cpr_inv_sort1: ∀L,T2,k. L ⊢ ⋆k ➡ T2 → T2 = ⋆k.
 >(tpss_inv_sort1 … H) -H //
 qed-.
 
+(* Basic_1: was pr2_gen_abbr *)
+lemma cpr_inv_abbr1: ∀L,V1,T1,U2. L ⊢ ⓓV1. T1 ➡ U2 →
+                     (∃∃V,V2,T2. V1 ➡ V & L ⊢ V [O, |L|] ▶* V2 &
+                                 L. ⓓV ⊢ T1 ➡ T2 &
+                                 U2 = ⓓV2. T2
+                      ) ∨
+                      ∃∃T. ⇧[0,1] T ≡ T1 & L ⊢ T ➡ U2.
+#L #V1 #T1 #Y * #X #H1 #H2
+elim (tpr_inv_abbr1 … H1) -H1 *
+[ #V #T0 #T #HV1 #HT10 #HT0 #H destruct
+  elim (tpss_inv_bind1 … H2) -H2 #V2 #T2 #HV2 #HT2 #H destruct
+  lapply (tps_lsubs_conf … HT0 (L. ⓓV) ?) -HT0 /2 width=1/ #HT0
+  lapply (tps_weak_all … HT0) -HT0 #HT0
+  lapply (tpss_lsubs_conf … HT2 (L. ⓓV) ?) -HT2 /2 width=1/ #HT2
+  lapply (tpss_weak_all … HT2) -HT2 #HT2
+  lapply (tpss_strap … HT0 HT2) -T /4 width=7/
+| /4 width=5/
+]
+qed-.
+
 (* Basic_1: was: pr2_gen_cast *)
-lemma cpr_inv_cast1: ∀L,V1,T1,U2. L ⊢ 𝕔{Cast} V1. T1 ➡ U2 → (
+lemma cpr_inv_cast1: ∀L,V1,T1,U2. L ⊢ V1. T1 ➡ U2 → (
                         ∃∃V2,T2. L ⊢ V1 ➡ V2 & L ⊢ T1 ➡ T2 &
-                                 U2 = 𝕔{Cast} V2. T2
+                                 U2 = V2. T2
                      ) ∨ L ⊢ T1 ➡ U2.
 #L #V1 #T1 #U2 * #X #H #HU2
 elim (tpr_inv_cast1 … H) -H /3 width=3/
@@ -84,13 +107,8 @@ elim (tpr_inv_cast1 … H) -H /3 width=3/
 elim (tpss_inv_flat1 … HU2) -HU2 #V2 #T2 #HV2 #HT2 #H destruct /4 width=5/
 qed-.
 
-(* Basic_1: removed theorems 5
-            pr2_head_1 pr2_head_2 pr2_cflat pr2_gen_cflat clear_pr2_trans
+(* Basic_1: removed theorems 4
+            pr2_head_2 pr2_cflat pr2_gen_cflat clear_pr2_trans
    Basic_1: removed local theorems 3:
             pr2_free_free pr2_free_delta pr2_delta_delta
 *)
-
-(*
-pr2/fwd pr2_gen_appl
-pr2/fwd pr2_gen_abbr
-*)