]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/lib/lambda-delta/reduction/pr_defs.ma
confluence of reduction started ...
[helm.git] / matita / matita / lib / lambda-delta / reduction / pr_defs.ma
index e1cd216c2277f253d070b1bce773618e94b101d7..d45cbd0d84393512a72683057868b6f2eff38a3d 100644 (file)
@@ -9,7 +9,7 @@
      \ /
       V_______________________________________________________________ *)
 
-include "lambda-delta/substitution/thin_defs.ma".
+include "lambda-delta/substitution/drop_defs.ma".
 
 (* SINGLE STEP PARALLEL REDUCTION ON TERMS **********************************)
 
@@ -24,7 +24,7 @@ inductive pr: lenv → term → term → Prop ≝
             pr L V1 V2 → pr (L. 𝕓{Abst} W) T1 T2 → (*𝕓*)
             pr L (𝕚{Appl} V1. 𝕚{Abst} W. T1) (𝕚{Abbr} V2. T2)
 | pr_delta: ∀L,K,V1,V2,V,i.
-            â\86\93[0,i] L â\89¡ K. ð\9d\95\93{Abbr} V1 → pr K V1 V2 → ↑[0,i+1] V2 ≡ V →
+            â\86\91[0,i] K. ð\9d\95\93{Abbr} V1 â\89¡ L → pr K V1 V2 → ↑[0,i+1] V2 ≡ V →
             pr L (#i) V
 | pr_theta: ∀L,V,V1,V2,W1,W2,T1,T2.
             pr L V1 V2 → ↑[0,1] V2 ≡ V → pr L W1 W2 → pr (L. 𝕓{Abbr} W1) T1 T2 → (*𝕓*)
@@ -38,22 +38,44 @@ interpretation
    "single step parallel reduction (term)"
    'PR L T1 T2 = (pr L T1 T2).
 
-(* The basic properties *****************************************************)
+(* Basic properties *********************************************************)
 
 lemma pr_refl: ∀T,L. L ⊢ T ⇒ T.
 #T elim T -T //
 #I elim I -I /2/
 qed.
-(*
-lemma subst_pr: ∀d,e,L,T1,U2. L ⊢ ↓[d,e] T1 ≡ U2 → ∀T2. ↑[d,e] U2 ≡ T2 →
-                L ⊢ T1 ⇒ T2.
-#d #e #L #T1 #U2 #H elim H -H d e L T1 U2
-[ #L #k #d #e #X #HX lapply (lift_inv_sort1 … HX) -HX #HX destruct -X // 
-| #L #i #d #e #Hid #X #HX lapply (lift_inv_sort1 … HX) -HX #HX destruct -X //
-| #L #V1 #V2 #e #HV12 * #V #HV2 #HV1
-  elim (lift_total 0 1 V1) #W1 #HVW1
-  @(ex2_1_intro … W1)
-  [
-  | /2 width=6/  
-
-*)
\ No newline at end of file
+
+(* The basic inversion lemmas ***********************************************)
+
+lemma pr_inv_lref2_aux: ∀L,T1,T2. L ⊢ T1 ⇒ T2 → ∀i. T2 = #i →
+                        ∨∨           T1 = #i
+                         | ∃∃K,V1,j. j < i & K ⊢ V1 ⇒ #(i-j-1) &
+                                     ↑[O,j] K. 𝕓{Abbr} V1 ≡ L & T1 = #j
+                         | ∃∃V,T,T0. ↑[O,1] T0 ≡ T & L ⊢ T0 ⇒ #i &
+                                     T1 = 𝕓{Abbr} V. T
+                         | ∃∃V,T.    L ⊢ T ⇒ #i & T1 = 𝕗{Cast} V. T.
+#L #T1 #T2 #H elim H -H L T1 T2
+[ #L #k #i #H destruct
+| #L #j #i /2/
+| #L #I #V1 #V2 #T1 #T2 #_ #_ #_ #_ #i #H destruct
+| #L #I #V1 #V2 #T1 #T2 #_ #_ #_ #_ #i #H destruct
+| #L #V1 #V2 #W #T1 #T2 #_ #_ #_ #_ #i #H destruct
+| #L #K #V1 #V2 #V #j #HLK #HV12 #HV2 #_ #i #H destruct -V;
+  elim (lift_inv_lref2 … HV2) -HV2 * #H1 #H2
+  [ elim (lt_zero_false … H1)
+  | destruct -V2 /3 width=7/
+  ]
+| #L #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #_ #_ #_ #i #H destruct
+| #L #V #T #T1 #T2 #HT1 #HT12 #_ #i #H destruct /3 width=6/
+| #L #V #T1 #T2 #HT12 #_ #i #H destruct /3/
+]
+qed.
+
+lemma pr_inv_lref2: ∀L,T1,i. L ⊢ T1 ⇒ #i →
+                    ∨∨           T1 = #i
+                     | ∃∃K,V1,j. j < i & K ⊢ V1 ⇒ #(i-j-1) &
+                                 ↑[O,j] K. 𝕓{Abbr} V1 ≡ L & T1 = #j
+                     | ∃∃V,T,T0. ↑[O,1] T0 ≡ T & L ⊢ T0 ⇒ #i &
+                                 T1 = 𝕓{Abbr} V. T
+                     | ∃∃V,T.    L ⊢ T ⇒ #i & T1 = 𝕗{Cast} V. T.
+/2/ qed.