]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/static/lfdeq.ma
notational update in ground_2 and basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / static / lfdeq.ma
index 8a770328f0758d0c063c21648c4058cabbb20915..47c5a38bad397caf4b5467d31d24b87b76853921 100644 (file)
@@ -28,14 +28,11 @@ interpretation
 interpretation
    "degree-based ranged equivalence (local environment)"
    'StarEqSn h o f L1 L2 = (lexs (cdeq_ext h o) cfull f L1 L2).
-(*
-definition lfdeq_transitive: predicate (relation3 lenv term term) ≝
-           λR. ∀L2,T1,T2. R L2 T1 T2 → ∀L1. L1 ≛[h, o, T1] L2 → R L1 T1 T2.
-*)
+
 (* Basic properties ***********************************************************)
 
-lemma frees_tdeq_conf_lexs: ∀h,o,f,L1,T1. L1 ⊢ 𝐅*⦃T1⦄ ≡ f → ∀T2. T1 ≛[h, o] T2 →
-                            ∀L2. L1 ≛[h, o, f] L2 → L2 ⊢ 𝐅*⦃T2⦄ ≡ f.
+lemma frees_tdeq_conf_lfdeq: ∀h,o,f,L1,T1. L1 ⊢ 𝐅*⦃T1⦄ ≘ f → ∀T2. T1 ≛[h, o] T2 →
+                             ∀L2. L1 ≛[h, o, f] L2 → L2 ⊢ 𝐅*⦃T2⦄ ≘ f.
 #h #o #f #L1 #T1 #H elim H -f -L1 -T1
 [ #f #L1 #s1 #Hf #X #H1 #L2 #_
   elim (tdeq_inv_sort1 … H1) -H1 #s2 #d #_ #_ #H destruct
@@ -68,27 +65,29 @@ lemma frees_tdeq_conf_lexs: ∀h,o,f,L1,T1. L1 ⊢ 𝐅*⦃T1⦄ ≡ f → ∀T2
 ]
 qed-.
 
-lemma frees_tdeq_conf: â\88\80h,o,f,L,T1. L â\8a¢ ð\9d\90\85*â¦\83T1â¦\84 â\89¡ f →
-                       â\88\80T2. T1 â\89\9b[h, o] T2 â\86\92 L â\8a¢ ð\9d\90\85*â¦\83T2â¦\84 â\89¡ f.
-/4 width=7 by frees_tdeq_conf_lexs, lexs_refl, ext2_refl/ qed-.
+lemma frees_tdeq_conf: â\88\80h,o,f,L,T1. L â\8a¢ ð\9d\90\85*â¦\83T1â¦\84 â\89\98 f →
+                       â\88\80T2. T1 â\89\9b[h, o] T2 â\86\92 L â\8a¢ ð\9d\90\85*â¦\83T2â¦\84 â\89\98 f.
+/4 width=7 by frees_tdeq_conf_lfdeq, lexs_refl, ext2_refl/ qed-.
 
-lemma frees_lexs_conf: ∀h,o,f,L1,T. L1 ⊢ 𝐅*⦃T⦄ ≡ f →
-                       ∀L2. L1 ≛[h, o, f] L2 → L2 ⊢ 𝐅*⦃T⦄ ≡ f.
-/2 width=7 by frees_tdeq_conf_lexs, tdeq_refl/ qed-.
+lemma frees_lfdeq_conf: ∀h,o,f,L1,T. L1 ⊢ 𝐅*⦃T⦄ ≘ f →
+                        ∀L2. L1 ≛[h, o, f] L2 → L2 ⊢ 𝐅*⦃T⦄ ≘ f.
+/2 width=7 by frees_tdeq_conf_lfdeq, tdeq_refl/ qed-.
 
-lemma frees_lfdeq_conf_lexs: ∀h,o. lexs_frees_confluent (cdeq_ext h o) cfull.
-/3 width=7 by frees_tdeq_conf_lexs, ex2_intro/ qed-.
-
-lemma tdeq_lfdeq_conf_sn: ∀h,o. s_r_confluent1 … (cdeq h o) (lfdeq h o).
-#h #o #L1 #T1 #T2 #HT12 #L2 *
+lemma tdeq_lfxs_conf: ∀R,h,o. s_r_confluent1 … (cdeq h o) (lfxs R).
+#R #h #o #L1 #T1 #T2 #HT12 #L2 *
 /3 width=5 by frees_tdeq_conf, ex2_intro/
 qed-.
 
-(* Basic_2A1: uses: lleq_sym *)
-lemma lfdeq_sym: ∀h,o,T. symmetric … (lfdeq h o T).
-#h #o #T #L1 #L2 *
-/4 width=7 by frees_tdeq_conf_lexs, lfxs_sym, tdeq_sym, ex2_intro/
-qed-.
+lemma tdeq_lfxs_div: ∀R,h,o,T1,T2. T1 ≛[h, o] T2 →
+                     ∀L1,L2. L1 ⪤*[R, T2] L2 → L1 ⪤*[R, T1] L2.
+/3 width=5 by tdeq_lfxs_conf, tdeq_sym/ qed-.
+
+lemma tdeq_lfdeq_conf: ∀h,o. s_r_confluent1 … (cdeq h o) (lfdeq h o).
+/2 width=5 by tdeq_lfxs_conf/ qed-.
+
+lemma tdeq_lfdeq_div: ∀h,o,T1,T2. T1 ≛[h, o] T2 →
+                      ∀L1,L2. L1 ≛[h, o, T2] L2 → L1 ≛[h, o, T1] L2.
+/2 width=5 by tdeq_lfxs_div/ qed-.
 
 lemma lfdeq_atom: ∀h,o,I. ⋆ ≛[h, o, ⓪{I}] ⋆.
 /2 width=1 by lfxs_atom/ qed.
@@ -107,7 +106,7 @@ lemma lfdeq_unit: ∀h,o,f,I,L1,L2. 𝐈⦃f⦄ → L1 ⪤*[cdeq_ext h o, cfull,
 /2 width=3 by lfxs_unit/ qed.
 *)
 lemma lfdeq_lref: ∀h,o,I1,I2,L1,L2,i.
-                  L1 â\89\9b[h, o, #i] L2 â\86\92 L1.â\93\98{I1} â\89\9b[h, o, #⫯i] L2.ⓘ{I2}.
+                  L1 â\89\9b[h, o, #i] L2 â\86\92 L1.â\93\98{I1} â\89\9b[h, o, #â\86\91i] L2.ⓘ{I2}.
 /2 width=1 by lfxs_lref/ qed.
 
 (* Basic_2A1: uses: lleq_gref *)
@@ -139,7 +138,7 @@ lemma lfdeq_inv_zero: ∀h,o,Y1,Y2. Y1 ≛[h, o, #0] Y2 →
 /3 width=9 by or3_intro0, or3_intro1, or3_intro2, ex4_5_intro, ex4_4_intro, conj/
 qed-.
 *)
-lemma lfdeq_inv_lref: â\88\80h,o,Y1,Y2,i. Y1 â\89\9b[h, o, #⫯i] Y2 →
+lemma lfdeq_inv_lref: â\88\80h,o,Y1,Y2,i. Y1 â\89\9b[h, o, #â\86\91i] Y2 →
                       (Y1 = ⋆ ∧ Y2 = ⋆) ∨
                       ∃∃I1,I2,L1,L2. L1 ≛[h, o, #i] L2 &
                                      Y1 = L1.ⓘ{I1} & Y2 = L2.ⓘ{I2}.
@@ -165,11 +164,11 @@ lemma lfdeq_inv_zero_pair_dx: ∀h,o,I,Y1,L2,V2. Y1 ≛[h, o, #0] L2.ⓑ{I}V2 
                               ∃∃L1,V1. L1 ≛[h, o, V1] L2 & V1 ≛[h, o] V2 & Y1 = L1.ⓑ{I}V1.
 /2 width=1 by lfxs_inv_zero_pair_dx/ qed-.
 
-lemma lfdeq_inv_lref_bind_sn: â\88\80h,o,I1,Y2,L1,i. L1.â\93\98{I1} â\89\9b[h, o, #⫯i] Y2 →
+lemma lfdeq_inv_lref_bind_sn: â\88\80h,o,I1,Y2,L1,i. L1.â\93\98{I1} â\89\9b[h, o, #â\86\91i] Y2 →
                               ∃∃I2,L2. L1 ≛[h, o, #i] L2 & Y2 = L2.ⓘ{I2}.
 /2 width=2 by lfxs_inv_lref_bind_sn/ qed-.
 
-lemma lfdeq_inv_lref_bind_dx: â\88\80h,o,I2,Y1,L2,i. Y1 â\89\9b[h, o, #⫯i] L2.ⓘ{I2} →
+lemma lfdeq_inv_lref_bind_dx: â\88\80h,o,I2,Y1,L2,i. Y1 â\89\9b[h, o, #â\86\91i] L2.ⓘ{I2} →
                               ∃∃I1,L1. L1 ≛[h, o, #i] L2 & Y1 = L1.ⓘ{I1}.
 /2 width=2 by lfxs_inv_lref_bind_dx/ qed-.
 
@@ -195,11 +194,3 @@ lemma lfdeq_fwd_flat_dx: ∀h,o,I,L1,L2,V,T. L1 ≛[h, o, ⓕ{I}V.T] L2 → L1 
 lemma lfdeq_fwd_dx: ∀h,o,I2,L1,K2. ∀T:term. L1 ≛[h, o, T] K2.ⓘ{I2} →
                     ∃∃I1,K1. L1 = K1.ⓘ{I1}.
 /2 width=5 by lfxs_fwd_dx/ qed-.
-
-(* Basic_2A1: removed theorems 10:
-              lleq_ind lleq_fwd_lref
-              lleq_fwd_drop_sn lleq_fwd_drop_dx
-              lleq_skip lleq_lref lleq_free
-              lleq_Y lleq_ge_up lleq_ge
-               
-*)