]> matita.cs.unibo.it Git - helm.git/commitdiff
- strong normalization of abbreviation proved
authorFerruccio Guidi <ferruccio.guidi@unibo.it>
Sat, 11 Feb 2012 19:48:08 +0000 (19:48 +0000)
committerFerruccio Guidi <ferruccio.guidi@unibo.it>
Sat, 11 Feb 2012 19:48:08 +0000 (19:48 +0000)
- bug fix in the generation lemma for abbreviationof context=sensitive parallel reduction

matita/matita/contribs/lambda_delta/Basic_2/computation/cprs.ma
matita/matita/contribs/lambda_delta/Basic_2/computation/csn.ma
matita/matita/contribs/lambda_delta/Basic_2/computation/csn_cprs.ma [new file with mode: 0644]
matita/matita/contribs/lambda_delta/Basic_2/computation/csn_lcpr.ma [new file with mode: 0644]
matita/matita/contribs/lambda_delta/Basic_2/reducibility/cpr.ma
matita/matita/contribs/lambda_delta/Basic_2/reducibility/cpr_ltpss.ma
matita/matita/contribs/lambda_delta/Basic_2/reducibility/lcpr_cpr.ma [new file with mode: 0644]

index 276abd46892c2af49b8f6cf520ddf6196d673af8..8748b6c3dedd390958dfa74958ea24778dc7ce63 100644 (file)
@@ -31,6 +31,10 @@ lemma cprs_ind: ∀L,T1. ∀R:predicate term. R T1 →
 #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 →
+                   (∀T1,T. L ⊢ T1 ➡ T → L ⊢ T ➡* T2 → R T → R T1) →
+                   ∀T1. L ⊢ T1 ➡* T2 → R T1.
+
 (* Basic properties *********************************************************)
 
 (* Basic_1: was: pr3_refl *)
index 2df67c9c226c7abfccbd81b1faef3f1ca0c0ae7a..b4106047214d2b20ed303e4280cdaa7564d552d2 100644 (file)
@@ -56,11 +56,6 @@ elim (term_eq_dec T1 T2) #HT12
 | -HT1 -HT2 /3 width=4/
 qed.
 
-axiom tpss_csn_trans: ∀L,T2. L ⊢ ⬇* T2 → ∀T1,d,e. L ⊢ T1 [d, e] ▶* T2 → L ⊢ ⬇* T1.
-(*
-#L #T2 #H @(csn_ind … H) -T2 #T2 #HT2 #IHT2 #T1 #d #e #HT12
-@csn_intro #T #HLT1 #HT1  
-*)
 (* Basic_1: was: sn3_cast *)
 lemma csn_cast: ∀L,W. L ⊢ ⬇* W → ∀T. L ⊢ ⬇* T → L ⊢ ⬇* ⓣW. T.
 #L #W #HW elim HW -W #W #_ #IHW #T #HT @(csn_ind … HT) -T #T #HT #IHT
diff --git a/matita/matita/contribs/lambda_delta/Basic_2/computation/csn_cprs.ma b/matita/matita/contribs/lambda_delta/Basic_2/computation/csn_cprs.ma
new file mode 100644 (file)
index 0000000..ace3b9f
--- /dev/null
@@ -0,0 +1,100 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+include "Basic_2/computation/cprs.ma".
+include "Basic_2/computation/csn.ma".
+
+(* CONTEXT-SENSITIVE STRONGLY NORMALIZING TERMS *****************************)
+
+(* Properties concerning context-sensitive computation on terms *************)
+
+definition csns: lenv → predicate term ≝ λL. SN … (cprs L) (eq …).
+
+interpretation
+   "context-sensitive strong normalization (term)"
+   'SNStar L T = (csns L T).
+
+notation "hvbox( L ⊢ ⬇ * * T )"
+   non associative with precedence 45
+   for @{ 'SNStar $L $T }.
+
+(* Basic eliminators ********************************************************)
+
+lemma csns_ind: ∀L. ∀R:predicate term.
+                (∀T1. L ⊢ ⬇** T1 →
+                      (∀T2. L ⊢ T1 ➡* T2 → (T1 = T2 → False) → R T2) → R T1
+                ) →
+                ∀T. L ⊢ ⬇** T → R T.
+#L #R #H0 #T1 #H elim H -T1 #T1 #HT1 #IHT1
+@H0 -H0 /3 width=1/ -IHT1 /4 width=1/
+qed-.
+
+(* Basic properties *********************************************************)
+
+lemma csns_intro: ∀L,T1.
+                  (∀T2. L ⊢ T1 ➡* T2 → (T1 = T2 → False) → L ⊢ ⬇** T2) → L ⊢ ⬇** T1.
+#L #T1 #H
+@(SN_intro … H)
+qed.
+
+fact csns_intro_aux: ∀L,T1.
+                     (∀T,T2. L ⊢ T ➡* T2 → T1 = T → (T1 = T2 → False) → L ⊢ ⬇** T2) → L ⊢ ⬇** T1.
+/4 width=3/ qed-.
+
+lemma csns_cprs_trans: ∀L,T1. L ⊢ ⬇** T1 → ∀T2. L ⊢ T1 ➡* T2 → L ⊢ ⬇** T2.
+#L #T1 #H elim H -T1 #T1 #HT1 #IHT1 #T2 #HLT12
+@csns_intro #T #HLT2 #HT2
+elim (term_eq_dec T1 T2) #HT12
+[ -IHT1 -HLT12 destruct /3 width=1/
+| -HT1 -HT2 /3 width=4/
+qed.
+
+lemma csns_intro_cpr: ∀L,T1.
+                      (∀T2. L ⊢ T1 ➡ T2 → (T1 = T2 → False) → L ⊢ ⬇** T2) →
+                      L ⊢ ⬇** T1.
+#L #T1 #H
+@csns_intro_aux #T #T2 #H @(cprs_ind_dx … H) -T
+[ -H #H destruct #H
+  elim (H ?) //
+| #T0 #T #HLT1 #HLT2 #IHT #HT10 #HT12 destruct
+  elim (term_eq_dec T0 T) #HT0
+  [ -HLT1 -HLT2 -H /3 width=1/
+  | -IHT -HT12 /4 width=3/
+  ]
+] 
+qed.
+
+(* Main properties **********************************************************)
+
+theorem csn_csns: ∀L,T. L ⊢ ⬇* T → L ⊢ ⬇** T.
+#L #T #H @(csn_ind … H) -T /4 width=1/
+qed.
+
+theorem csns_csn: ∀L,T. L ⊢ ⬇** T → L ⊢ ⬇* T.
+#L #T #H @(csns_ind … H) -T /4 width=1/
+qed.
+
+lemma csn_cprs_trans: ∀L,T1. L ⊢ ⬇* T1 → ∀T2. L ⊢ T1 ➡* T2 → L ⊢ ⬇* T2.
+/4 width=3/ qed.
+
+(* Main eliminators *********************************************************)
+
+lemma csn_ind_cprs: ∀L. ∀R:predicate term.
+                    (∀T1. L ⊢ ⬇* T1 →
+                          (∀T2. L ⊢ T1 ➡* T2 → (T1 = T2 → False) → R T2) → R T1
+                    ) →
+                    ∀T. L ⊢ ⬇* T → R T.
+#L #R #H0 #T1 #H @(csns_ind … (csn_csns … H)) -T1 #T1 #HT1 #IHT1
+@H0 -H0 /2 width=1/ -HT1 /3 width=1/
+qed-.
diff --git a/matita/matita/contribs/lambda_delta/Basic_2/computation/csn_lcpr.ma b/matita/matita/contribs/lambda_delta/Basic_2/computation/csn_lcpr.ma
new file mode 100644 (file)
index 0000000..88c4927
--- /dev/null
@@ -0,0 +1,45 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+include "Basic_2/reducibility/lcpr_cpr.ma".
+include "Basic_2/computation/cprs_lcpr.ma".
+include "Basic_2/computation/csn_cprs.ma".
+include "Basic_2/computation/csn_lift.ma".
+
+(* CONTEXT-SENSITIVE STRONGLY NORMALIZING TERMS *****************************)
+
+(* Advanced properties ******************************************************)
+
+lemma csn_lcpr_trans: ∀L1,L2. L1 ⊢ ➡ L2 → ∀T. L1 ⊢ ⬇* T → L2 ⊢ ⬇* T.
+#L1 #L2 #HL12 #T #H @(csn_ind_cprs … H) -T #T #_ #IHT
+@csn_intro #T0 #HLT0 #HT0
+@IHT /2 width=2/ -IHT -HT0 /2 width=3/
+qed.
+
+lemma csn_abbr: ∀L,V. L ⊢ ⬇* V → ∀T. L. ⓓV ⊢ ⬇* T → L ⊢ ⬇* ⓓV. T.
+#L #V #HV @(csn_ind … HV) -V #V #_ #IHV #T #HT @(csn_ind_cprs … HT) -T #T #HT #IHT
+@csn_intro #X #H1 #H2
+elim (cpr_inv_abbr1 … H1) -H1 *
+[ #V0 #V1 #T1 #HLV0 #HLV01 #HLT1 #H destruct
+  lapply (cpr_intro … HLV0 HLV01) -HLV01 #HLV1
+  lapply (ltpr_cpr_trans (L. ⓓV) … HLT1) /2 width=1/ -V0 #HLT1
+  elim (eq_false_inv_tpair … H2) -H2
+  [ #HV1 @IHV // /2 width=1/ -HV1
+    @(csn_lcpr_trans (L. ⓓV)) /2 width=1/ -HLV1 /2 width=3/
+  | -IHV -HLV1 * #H destruct /3 width=1/
+  ]
+| -IHV -IHT -H2 #T0 #HT0 #HLT0
+  lapply (csn_inv_lift … HT … HT0) -HT /2 width=3/
+]
+qed.
index 4edfc6456969909b09b831c72b67a8a74ec1a736..fca70da65141863270ffc4f2c13035995c5dd143 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.
@@ -73,6 +76,26 @@ 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 ⊢ ⓣV1. T1 ➡ U2 → (
                         ∃∃V2,T2. L ⊢ V1 ➡ V2 & L ⊢ T1 ➡ T2 &
index d925d19c5e7033dba3627c94aebf3bd3ef1f4b65..148a29645c81cd5c017aba2fe07406f0c6c4d50f 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-include "Basic_2/unfold/ltpss_ltpss.ma".
+include "Basic_2/unfold/ltpss_tpss.ma".
 include "Basic_2/reducibility/cpr.ma".
 
 (* CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS ****************************)
 
-(* Advanced inversion lemmas ************************************************)
-
-(* Basic_1: was pr2_gen_abbr *)
-lemma cpr_inv_abbr1: ∀L,V1,T1,U2. L ⊢ ⓓV1. T1 ➡ U2 →
-                     (∃∃V2,T2,T. L ⊢ V1 ➡ V2 & L. ⓓV2 ⊢ T1 ➡ T2 &
-                                 L.  ⓓV2 ⊢ T [1, |L|] ▶* T2 &
-                                 U2 = ⓓV2. T
-                      ) ∨
-                      ∃∃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
-  elim (ltpss_tps_conf … HT0 (L. ⓓV2) 1 (|L|) ?) -HT0 /2 width=1/ #V0 #HV20 #HV0
-  lapply (tpss_weak_all … HV20) -HV20 #HV20
-  lapply (tpss_lsubs_conf … HV0 (L. ⓓV2) ?) -HV0 /2 width=1/ #HV0
-  elim (tpss_conf_eq … HT2 … HV0) -T #T3 #HT23 #HVT3
-  lapply (tpss_weak_all … HVT3) -HVT3 #HVT3
-  lapply (tpss_trans_eq … HV20 … HVT3) -V0 /4 width=7/
-| /4 width=5/
-]
-qed-.
-
 (* Properties concerning partial unfold on local environments ***************)
 
 lemma ltpss_cpr_trans: ∀L1,L2,d,e. L1 [d, e] ▶* L2 →
diff --git a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/lcpr_cpr.ma b/matita/matita/contribs/lambda_delta/Basic_2/reducibility/lcpr_cpr.ma
new file mode 100644 (file)
index 0000000..7abe347
--- /dev/null
@@ -0,0 +1,27 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+include "Basic_2/unfold/ltpss_ltpss.ma".
+include "Basic_2/reducibility/cpr.ma".
+include "Basic_2/reducibility/lcpr.ma".
+
+(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON LOCAL ENVIRONMENTS *************)
+
+(* Advanced properties ****************************************************)
+
+lemma lcpr_pair: ∀L1,L2. L1 ⊢ ➡ L2 → ∀V1,V2. L2 ⊢ V1 ➡ V2 →
+                 ∀I. L1. ⓑ{I} V1 ⊢ ➡ L2. ⓑ{I} V2.
+#L1 #L2 * #L #HL1 #HL2 #V1 #V2 *
+<(ltpss_fwd_length … HL2) /4 width=5/
+qed.