]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/Basic_2/computation/csn_lift.ma
- property S6 of stronfly normalizing terms proved
[helm.git] / matita / matita / contribs / lambda_delta / Basic_2 / computation / csn_lift.ma
index 29872f92903bc01b67aad87a15967554a6154ec8..01d983c0cd0435af8e263b772932f77738fdcf92 100644 (file)
@@ -43,15 +43,6 @@ qed.
 
 (* Advanced properties ******************************************************)
 
-lemma csn_acp: acp cpr (eq …) (csn …).
-@mk_acp
-[ /2 width=1/
-| /2 width=3/
-| /2 width=5/
-| @cnf_lift
-]
-qed.
-
 (* Basic_1: was: sn3_abbr *)
 lemma csn_lref_abbr: ∀L,K,V,i. ⇩[0, i] L ≡ K. ⓓV → K ⊢ ⬇* V → L ⊢ ⬇* #i.
 #L #K #V #i #HLK #HV
@@ -72,8 +63,38 @@ lemma csn_abst: ∀L,W. L ⊢ ⬇* W → ∀I,V,T. L. ⓑ{I} V ⊢ ⬇* T → L
 @csn_intro #X #H1 #H2
 elim (cpr_inv_abst1 … H1 I V) -H1
 #W0 #T0 #HLW0 #HLT0 #H destruct
-elim (eq_false_inv_tpair … H2) -H2
+elim (eq_false_inv_tpair_sn … H2) -H2
 [ /3 width=5/
 | -HLW0 * #H destruct /3 width=1/
 ]
 qed.
+
+lemma csn_appl_simple: ∀L,V. L ⊢ ⬇* V → ∀T1.
+                       (∀T2. L ⊢ T1 ➡ T2 → (T1 = T2 → False) → L ⊢ ⬇* ⓐV. T2) →
+                       𝐒[T1] → L ⊢ ⬇* ⓐV. T1.
+#L #V #H @(csn_ind … H) -V #V #_ #IHV #T1 #IHT1 #HT1
+@csn_intro #X #H1 #H2
+elim (cpr_inv_appl1_simple … H1 ?) // -H1
+#V0 #T0 #HLV0 #HLT10 #H destruct
+elim (eq_false_inv_tpair_dx … H2) -H2
+[ -IHV -HT1 #HT10
+  @(csn_cpr_trans … (ⓐV.T0)) /2 width=1/ -HLV0
+  @IHT1 -IHT1 // /2 width=1/
+| -HLT10 * #H #HV0 destruct
+  @IHV -IHV // -HT1 /2 width=1/ -HV0
+  #T2 #HLT02 #HT02
+  @(csn_cpr_trans … (ⓐV.T2)) /2 width=1/ -HLV0
+  @IHT1 -IHT1 // -HLT02 /2 width=1/
+]
+qed.
+
+(* Main properties **********************************************************)
+
+theorem csn_acp: acp cpr (eq …) (csn …).
+@mk_acp
+[ /2 width=1/
+| /2 width=3/
+| /2 width=5/
+| @cnf_lift
+]
+qed.