]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/basic_2/computation/cprs.ma
- a caracterization of the top elements of the local evironment
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / computation / cprs.ma
index 4341a5a94a6d81219405893f98d49c9b21f7d85d..7bfe248954f64fee47f777666f48722852ac4ef8 100644 (file)
@@ -28,12 +28,16 @@ interpretation "context-sensitive parallel computation (term)"
 lemma cprs_ind: ∀L,T1. ∀R:predicate term. R T1 →
                 (∀T,T2. L ⊢ T1 ➡* T → L ⊢ T ➡ T2 → R T → R T2) →
                 ∀T2. L ⊢ T1 ➡* T2 → R T2.
-#L #T1 #R #HT1 #IHT1 #T2 #HT12 @(TC_star_ind … HT1 IHT1 … HT12) //
+#L #T1 #R #HT1 #IHT1 #T2 #HT12
+@(TC_star_ind … HT1 IHT1 … HT12) //
 qed-.
 
-axiom cprs_ind_dx: ∀L,T2. ∀R:predicate term. R T2 →
+lemma cprs_ind_dx: ∀L,T2. ∀R:predicate term. R T2 →
                    (∀T1,T. L ⊢ T1 ➡ T → L ⊢ T ➡* T2 → R T → R T1) →
                    ∀T1. L ⊢ T1 ➡* T2 → R T1.
+#L #T2 #R #HT2 #IHT2 #T1 #HT12
+@(TC_star_ind_dx … HT2 IHT2 … HT12) //
+qed-.
 
 (* Basic properties *********************************************************)
 
@@ -52,10 +56,11 @@ lemma cprs_strap2: ∀L,T1,T,T2.
 
 (* Note: it does not hold replacing |L1| with |L2| *)
 lemma cprs_lsubs_conf: ∀L1,T1,T2. L1 ⊢ T1 ➡* T2 →
-                       ∀L2. L1 [0, |L1|] ≼ L2 → L2 ⊢ T1 ➡* T2.
+                       ∀L2. L1 ≼ [0, |L1|] L2 → L2 ⊢ T1 ➡* T2.
 /3 width=3/
 qed.
 
+(* Basic_1: was only: pr3_thin_dx *)
 lemma cprs_flat_dx: ∀I,L,V1,V2. L ⊢ V1 ➡ V2 → ∀T1,T2. L ⊢ T1 ➡* T2 →
                     L ⊢ ⓕ{I} V1. T1 ➡* ⓕ{I} V2. T2.
 #I #L #V1 #V2 #HV12 #T1 #T2 #HT12 @(cprs_ind … HT12) -T2 /3 width=1/
@@ -82,13 +87,15 @@ elim (cpr_inv_cast1 … HU2) -HU2 /3 width=3/ *
 #W2 #T2 #HW2 #HT2 #H destruct /4 width=5/
 qed-.
 
+(* Basic_1: was: nf2_pr3_unfold *)
 lemma cprs_inv_cnf1: ∀L,T,U. L ⊢ T ➡* U → L ⊢ 𝐍[T] → T = U.
 #L #T #U #H @(cprs_ind_dx … H) -T //
 #T0 #T #H1T0 #_ #IHT #H2T0
 lapply (H2T0 … H1T0) -H1T0 #H destruct /2 width=1/
 qed-.
 
-(* Basic_1: removed theorems 6:
+(* Basic_1: removed theorems 10:
    clear_pr3_trans pr3_cflat pr3_gen_bind
+   pr3_head_1 pr3_head_2 pr3_head_21 pr3_head_12
    pr3_iso_appl_bind pr3_iso_appls_appl_bind pr3_iso_appls_bind
 *)