X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fetc%2Flpx_sn%2Flpx_sn_tc.etc;h=51b09380f28920c2f1ac6bca6e559d6e52bc3834;hb=db020b4218272e2e35641ce3bc3b0a9b3afda899;hp=2919d9d103b73f716e4f1c92c4110ad3f9bf7255;hpb=956df16197063a88b3858e3d212d7ed0f2c5ff46;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/lpx_sn/lpx_sn_tc.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/lpx_sn/lpx_sn_tc.etc index 2919d9d10..51b09380f 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/etc/lpx_sn/lpx_sn_tc.etc +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/lpx_sn/lpx_sn_tc.etc @@ -12,7 +12,7 @@ (* *) (**************************************************************************) -include "basic_2/grammar/lpx_sn.ma". +include "basic_2/substitution/lpx_sn.ma". (* SN POINTWISE EXTENSION OF A CONTEXT-SENSITIVE REALTION FOR TERMS *********) @@ -38,13 +38,6 @@ lemma TC_lpx_sn_pair: ∀R. (∀L. reflexive … (R L)) → ] qed-. -lemma lpx_sn_LTC_TC_lpx_sn: ∀R. (∀L. reflexive … (R L)) → - ∀L1,L2. lpx_sn (LTC … R) L1 L2 → - TC … (lpx_sn R) L1 L2. -#R #HR #L1 #L2 #H elim H -L1 -L2 -/2 width=1 by TC_lpx_sn_pair, lpx_sn_atom, inj/ -qed-. - (* Inversion lemmas on transitive closure ***********************************) lemma TC_lpx_sn_inv_atom2: ∀R,L1. TC … (lpx_sn R) L1 (⋆) → L1 = ⋆. @@ -54,7 +47,7 @@ lemma TC_lpx_sn_inv_atom2: ∀R,L1. TC … (lpx_sn R) L1 (⋆) → L1 = ⋆. ] qed-. -lemma TC_lpx_sn_inv_pair2: ∀R. s_rs_trans … R (lpx_sn R) → +lemma TC_lpx_sn_inv_pair2: ∀R. s_rs_transitive … R (λ_. lpx_sn R) → ∀I,L1,K2,V2. TC … (lpx_sn R) L1 (K2.ⓑ{I}V2) → ∃∃K1,V1. TC … (lpx_sn R) K1 K2 & LTC … R K1 V1 V2 & L1 = K1. ⓑ{I} V1. #R #HR #I #L1 #K2 #V2 #H @(TC_ind_dx … L1 H) -L1 @@ -65,22 +58,6 @@ lemma TC_lpx_sn_inv_pair2: ∀R. s_rs_trans … R (lpx_sn R) → ] qed-. -lemma TC_lpx_sn_ind: ∀R. s_rs_trans … R (lpx_sn R) → - ∀S:relation lenv. - S (⋆) (⋆) → ( - ∀I,K1,K2,V1,V2. - TC … (lpx_sn R) K1 K2 → LTC … R K1 V1 V2 → - S K1 K2 → S (K1.ⓑ{I}V1) (K2.ⓑ{I}V2) - ) → - ∀L2,L1. TC … (lpx_sn R) L1 L2 → S L1 L2. -#R #HR #S #IH1 #IH2 #L2 elim L2 -L2 -[ #X #H >(TC_lpx_sn_inv_atom2 … H) -X // -| #L2 #I #V2 #IHL2 #X #H - elim (TC_lpx_sn_inv_pair2 … H) // -H -HR - #L1 #V1 #HL12 #HV12 #H destruct /3 width=1 by/ -] -qed-. - lemma TC_lpx_sn_inv_atom1: ∀R,L2. TC … (lpx_sn R) (⋆) L2 → L2 = ⋆. #R #L2 #H elim H -L2 [ /2 width=2 by lpx_sn_inv_atom1/ @@ -88,7 +65,7 @@ lemma TC_lpx_sn_inv_atom1: ∀R,L2. TC … (lpx_sn R) (⋆) L2 → L2 = ⋆. ] qed-. -fact TC_lpx_sn_inv_pair1_aux: ∀R. s_rs_trans … R (lpx_sn R) → +fact TC_lpx_sn_inv_pair1_aux: ∀R. s_rs_transitive … R (λ_. lpx_sn R) → ∀L1,L2. TC … (lpx_sn R) L1 L2 → ∀I,K1,V1. L1 = K1.ⓑ{I}V1 → ∃∃K2,V2. TC … (lpx_sn R) K1 K2 & LTC … R K1 V1 V2 & L2 = K2. ⓑ{I} V2. @@ -98,16 +75,11 @@ fact TC_lpx_sn_inv_pair1_aux: ∀R. s_rs_trans … R (lpx_sn R) → ] qed-. -lemma TC_lpx_sn_inv_pair1: ∀R. s_rs_trans … R (lpx_sn R) → +lemma TC_lpx_sn_inv_pair1: ∀R. s_rs_transitive … R (λ_. lpx_sn R) → ∀I,K1,L2,V1. TC … (lpx_sn R) (K1.ⓑ{I}V1) L2 → ∃∃K2,V2. TC … (lpx_sn R) K1 K2 & LTC … R K1 V1 V2 & L2 = K2. ⓑ{I} V2. /2 width=3 by TC_lpx_sn_inv_pair1_aux/ qed-. -lemma TC_lpx_sn_inv_lpx_sn_LTC: ∀R. s_rs_trans … R (lpx_sn R) → - ∀L1,L2. TC … (lpx_sn R) L1 L2 → - lpx_sn (LTC … R) L1 L2. -/3 width=4 by TC_lpx_sn_ind, lpx_sn_pair/ qed-. - (* Forward lemmas on transitive closure *************************************) lemma TC_lpx_sn_fwd_length: ∀R,L1,L2. TC … (lpx_sn R) L1 L2 → |L1| = |L2|.