]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/static/lfdeq.ma
renaming
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / static / lfdeq.ma
index 6fb68863e758177595807af0401070bc8d233c7f..4103c54c50ce4e191e7a55ce9b876df1e8ee52c5 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/notation/relations/lazyeq_5.ma".
-include "basic_2/syntax/tdeq.ma".
+include "basic_2/notation/relations/stareqsn_5.ma".
+include "basic_2/syntax/tdeq_ext.ma".
 include "basic_2/static/lfxs.ma".
 
 (* DEGREE-BASED EQUIVALENCE FOR LOCAL ENVIRONMENTS ON REFERRED ENTRIES ******)
 
-definition lfdeq: ∀h. sd h → relation3 term lenv lenv ≝
-                  λh,o. lfxs (cdeq h o).
+definition lfdeq (h) (o): relation3 term lenv lenv ≝
+                          lfxs (cdeq h o).
 
 interpretation
    "degree-based equivalence on referred entries (local environment)"
-   'LazyEq h o T L1 L2 = (lfdeq h o T L1 L2).
+   'StarEqSn h o T L1 L2 = (lfdeq h o T L1 L2).
 
 interpretation
    "degree-based ranged equivalence (local environment)"
-   'LazyEq h o f L1 L2 = (lexs (cdeq 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.
-*)
+   'StarEqSn h o f L1 L2 = (lexs (cdeq_ext h o) cfull f L1 L2).
+
 (* 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 #I1 #Hf #X #H1 elim (tdeq_fwd_atom1 … H1) -H1
-  #I2 #H1 #Y #H2 lapply (lexs_inv_atom1 … H2) -H2
-  #H2 destruct /2 width=1 by frees_atom/
-| #f #I #L1 #V1 #s1 #_ #IH #X #H1 elim (tdeq_inv_sort1 … H1) -H1
-  #s2 #d #Hs1 #Hs2 #H1 #Y #H2 elim (lexs_inv_push1 … H2) -H2
-  #L2 #V2 #HL12 #_ #H2 destruct /4 width=3 by frees_sort, tdeq_sort/
-| #f #I #L1 #V1 #_ #IH #X #H1 >(tdeq_inv_lref1 … H1) -H1
-  #Y #H2 elim (lexs_inv_next1 … H2) -H2
-  #L2 #V2 #HL12 #HV12 #H2 destruct /3 width=1 by frees_zero/
-| #f #I #L1 #V1 #i #_ #IH #X #H1 >(tdeq_inv_lref1 … H1) -H1
-  #Y #H2 elim (lexs_inv_push1 … H2) -H2
-  #L2 #V2 #HL12 #_ #H2 destruct /3 width=1 by frees_lref/
-| #f #I #L1 #V1 #l #_ #IH #X #H1 >(tdeq_inv_gref1 … H1) -H1
-  #Y #H2 elim (lexs_inv_push1 … H2) -H2
-  #L2 #V2 #HL12 #_ #H2 destruct /3 width=1 by frees_gref/
-| #f1V #f1T #f1 #p #I #L1 #V1 #T1 #_ #_ #Hf1 #IHV #IHT #X #H1 elim (tdeq_inv_pair1 … H1) -H1
-  #V2 #T2 #HV12 #HT12 #H1 #L2 #HL12 destruct
-  /6 width=5 by frees_bind, lexs_inv_tl, sle_lexs_trans, sor_inv_sle_dx, sor_inv_sle_sn/
-| #f1V #f1T #f1 #I #L1 #V1 #T1 #_ #_ #Hf1 #IHV #IHT #X #H1 elim (tdeq_inv_pair1 … H1) -H1
-  #V2 #T2 #HV12 #HT12 #H1 #L2 #HL12 destruct
+[ #f #L1 #s1 #Hf #X #H1 #L2 #_
+  elim (tdeq_inv_sort1 … H1) -H1 #s2 #d #_ #_ #H destruct
+  /2 width=3 by frees_sort/
+| #f #i #Hf #X #H1
+  >(tdeq_inv_lref1 … H1) -X #Y #H2
+  >(lexs_inv_atom1 … H2) -Y
+  /2 width=1 by frees_atom/
+| #f #I #L1 #V1 #_ #IH #X #H1
+  >(tdeq_inv_lref1 … H1) -X #Y #H2
+  elim (lexs_inv_next1 … H2) -H2 #Z #L2 #HL12 #HZ #H destruct
+  elim (ext2_inv_pair_sn … HZ) -HZ #V2 #HV12 #H destruct
+  /3 width=1 by frees_pair/
+| #f #I #L1 #Hf #X #H1
+  >(tdeq_inv_lref1 … H1) -X #Y #H2
+  elim (lexs_inv_next1 … H2) -H2 #Z #L2 #_ #HZ #H destruct
+  >(ext2_inv_unit_sn … HZ) -Z /2 width=1 by frees_unit/
+| #f #I #L1 #i #_ #IH #X #H1
+  >(tdeq_inv_lref1 … H1) -X #Y #H2
+  elim (lexs_inv_push1 … H2) -H2 #J #L2 #HL12 #_ #H destruct
+  /3 width=1 by frees_lref/
+| #f #L1 #l #Hf #X #H1 #L2 #_
+  >(tdeq_inv_gref1 … H1) -X /2 width=1 by frees_gref/
+| #f1V #f1T #f1 #p #I #L1 #V1 #T1 #_ #_ #Hf1 #IHV #IHT #X #H1
+  elim (tdeq_inv_pair1 … H1) -H1 #V2 #T2 #HV12 #HT12 #H1 #L2 #HL12 destruct
+  /6 width=5 by frees_bind, lexs_inv_tl, ext2_pair, sle_lexs_trans, sor_inv_sle_dx, sor_inv_sle_sn/
+| #f1V #f1T #f1 #I #L1 #V1 #T1 #_ #_ #Hf1 #IHV #IHT #X #H1
+  elim (tdeq_inv_pair1 … H1) -H1 #V2 #T2 #HV12 #HT12 #H1 #L2 #HL12 destruct
   /5 width=5 by frees_flat, sle_lexs_trans, sor_inv_sle_dx, sor_inv_sle_sn/
 ]
 qed-.
 
-lemma frees_tdeq_conf: ∀h,o,f,L,T1. L ⊢ 𝐅*⦃T1⦄ ≡ f →
-                       ∀T2. T1 ≡[h, o] T2 → L ⊢ 𝐅*⦃T2⦄ ≡ f.
-/3 width=7 by frees_tdeq_conf_lexs, lexs_refl/ qed-.
+lemma frees_tdeq_conf (h) (o): ∀f,L,T1. L ⊢ 𝐅*⦃T1⦄ ≘ f →
+                               ∀T2. T1 ≛[h, o] T2 → L ⊢ 𝐅*⦃T2⦄ ≘ f.
+/4 width=7 by frees_tdeq_conf_lfdeq, lexs_refl, ext2_refl/ qed-.
 
-lemma frees_lfdeq_conf_lexs: ∀h,o. lexs_frees_confluent (cdeq h o) cfull.
-/3 width=7 by frees_tdeq_conf_lexs, ex2_intro/ 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 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 lfdeq_atom: ∀h,o,I. ⋆ ≡[h, o, ⓪{I}] ⋆.
+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.
 
-(* Basic_2A1: uses: lleq_sort *)
-lemma lfdeq_sort: ∀h,o,I,L1,L2,V1,V2,s.
-                  L1 ≡[h, o, ⋆s] L2 → L1.ⓑ{I}V1 ≡[h, o, ⋆s] L2.ⓑ{I}V2.
+lemma lfdeq_sort (h) (o): ∀I1,I2,L1,L2,s.
+                          L1 ≛[h, o, ⋆s] L2 → L1.ⓘ{I1} ≛[h, o, ⋆s] L2.ⓘ{I2}.
 /2 width=1 by lfxs_sort/ qed.
 
-lemma lfdeq_zero: ∀h,o,I,L1,L2,V.
-                  L1 ≡[h, o, V] L2 → L1.ⓑ{I}V ≡[h, o, #0] L2.ⓑ{I}V.
-/2 width=1 by lfxs_zero/ qed.
-
-lemma lfdeq_lref: ∀h,o,I,L1,L2,V1,V2,i.
-                  L1 ≡[h, o, #i] L2 → L1.ⓑ{I}V1 ≡[h, o, #⫯i] L2.ⓑ{I}V2.
+lemma lfdeq_pair (h) (o): ∀I,L1,L2,V1,V2. L1 ≛[h, o, V1] L2 → V1 ≛[h, o] V2 →
+                          L1.ⓑ{I}V1 ≛[h, o, #0] L2.ⓑ{I}V2.
+/2 width=1 by lfxs_pair/ qed.
+(*
+lemma lfdeq_unit (h) (o): ∀f,I,L1,L2. 𝐈⦃f⦄ → L1 ⪤*[cdeq_ext h o, cfull, f] L2 →
+                          L1.ⓤ{I} ≛[h, o, #0] L2.ⓤ{I}.
+/2 width=3 by lfxs_unit/ qed.
+*)
+lemma lfdeq_lref (h) (o): ∀I1,I2,L1,L2,i.
+                          L1 ≛[h, o, #i] L2 → L1.ⓘ{I1} ≛[h, o, #↑i] L2.ⓘ{I2}.
 /2 width=1 by lfxs_lref/ qed.
 
-(* Basic_2A1: uses: lleq_gref *)
-lemma lfdeq_gref: ∀h,o,I,L1,L2,V1,V2,l.
-                  L1 ≡[h, o, §l] L2 → L1.ⓑ{I}V1 ≡[h, o, §l] L2.ⓑ{I}V2.
+lemma lfdeq_gref (h) (o): ∀I1,I2,L1,L2,l.
+                          L1 ≛[h, o, §l] L2 → L1.ⓘ{I1} ≛[h, o, §l] L2.ⓘ{I2}.
 /2 width=1 by lfxs_gref/ qed.
 
-lemma lfdeq_pair_repl_dx: ∀h,o,I,L1,L2.∀T:term.∀V,V1.
-                          L1.ⓑ{I}V ≡[h, o, T] L2.ⓑ{I}V1 →
-                          ∀V2. V ≡[h, o] V2 →
-                          L1.ⓑ{I}V ≡[h, o, T] L2.ⓑ{I}V2.
-/2 width=2 by lfxs_pair_repl_dx/ qed-.
+lemma lfdeq_bind_repl_dx (h) (o): ∀I,I1,L1,L2.∀T:term.
+                                  L1.ⓘ{I} ≛[h, o, T] L2.ⓘ{I1} →
+                                  ∀I2. I ≛[h, o] I2 →
+                                  L1.ⓘ{I} ≛[h, o, T] L2.ⓘ{I2}.
+/2 width=2 by lfxs_bind_repl_dx/ qed-.
 
 (* Basic inversion lemmas ***************************************************)
 
-lemma lfdeq_inv_atom_sn: ∀h,o,Y2. ∀T:term. ⋆ ≡[h, o, T] Y2 → Y2 = ⋆.
+lemma lfdeq_inv_atom_sn (h) (o): ∀Y2. ∀T:term. ⋆ ≛[h, o, T] Y2 → Y2 = ⋆.
 /2 width=3 by lfxs_inv_atom_sn/ qed-.
 
-lemma lfdeq_inv_atom_dx: ∀h,o,Y1. ∀T:term. Y1 ≡[h, o, T] ⋆ → Y1 = ⋆.
+lemma lfdeq_inv_atom_dx (h) (o): ∀Y1. ∀T:term. Y1 ≛[h, o, T] ⋆ → Y1 = ⋆.
 /2 width=3 by lfxs_inv_atom_dx/ qed-.
-
-lemma lfdeq_inv_zero: ∀h,o,Y1,Y2. Y1 ≡[h, o, #0] Y2 →
-                      (Y1 = ⋆ ∧ Y2 = ⋆) ∨
-                      ∃∃I,L1,L2,V1,V2. L1 ≡[h, o, V1] L2 & V1 ≡[h, o] V2 &
-                                       Y1 = L1.ⓑ{I}V1 & Y2 = L2.ⓑ{I}V2.
+(*
+lemma lfdeq_inv_zero (h) (o): ∀Y1,Y2. Y1 ≛[h, o, #0] Y2 →
+                              ∨∨ ∧∧ Y1 = ⋆ & Y2 = ⋆
+                               | ∃∃I,L1,L2,V1,V2. L1 ≛[h, o, V1] L2 & V1 ≛[h, o] V2 &
+                                                  Y1 = L1.ⓑ{I}V1 & Y2 = L2.ⓑ{I}V2
+                               | ∃∃f,I,L1,L2. 𝐈⦃f⦄ & L1 ⪤*[cdeq_ext h o, cfull, f] L2 &
+                                              Y1 = L1.ⓤ{I} & Y2 = L2.ⓤ{I}.
 #h #o #Y1 #Y2 #H elim (lfxs_inv_zero … H) -H *
-/3 width=9 by ex4_5_intro, or_introl, or_intror, conj/
+/3 width=9 by or3_intro0, or3_intro1, or3_intro2, ex4_5_intro, ex4_4_intro, conj/
 qed-.
-
-lemma lfdeq_inv_lref: ∀h,o,Y1,Y2,i. Y1 ≡[h, o, #⫯i] Y2 →
-                      (Y1 = ⋆ ∧ Y2 = ⋆) ∨ 
-                      ∃∃I,L1,L2,V1,V2. L1 ≡[h, o, #i] L2 &
-                                       Y1 = L1.ⓑ{I}V1 & Y2 = L2.ⓑ{I}V2.
+*)
+lemma lfdeq_inv_lref (h) (o): ∀Y1,Y2,i. Y1 ≛[h, o, #↑i] Y2 →
+                              ∨∨ ∧∧ Y1 = ⋆ & Y2 = ⋆
+                               | ∃∃I1,I2,L1,L2. L1 ≛[h, o, #i] L2 &
+                                                Y1 = L1.ⓘ{I1} & Y2 = L2.ⓘ{I2}.
 /2 width=1 by lfxs_inv_lref/ qed-.
 
 (* Basic_2A1: uses: lleq_inv_bind lleq_inv_bind_O *)
-lemma lfdeq_inv_bind: ∀h,o,p,I,L1,L2,V,T. L1 ≡[h, o, ⓑ{p,I}V.T] L2 →
-                      L1 ≡[h, o, V] L2 ∧ L1.ⓑ{I}V ≡[h, o, T] L2.ⓑ{I}V.
+lemma lfdeq_inv_bind (h) (o): ∀p,I,L1,L2,V,T. L1 ≛[h, o, ⓑ{p,I}V.T] L2 →
+                              ∧∧ L1 ≛[h, o, V] L2 & L1.ⓑ{I}V ≛[h, o, T] L2.ⓑ{I}V.
 /2 width=2 by lfxs_inv_bind/ qed-.
 
 (* Basic_2A1: uses: lleq_inv_flat *)
-lemma lfdeq_inv_flat: ∀h,o,I,L1,L2,V,T. L1 ≡[h, o, ⓕ{I}V.T] L2 →
-                      L1 ≡[h, o, V] L2 ∧ L1 ≡[h, o, T] L2.
+lemma lfdeq_inv_flat (h) (o): ∀I,L1,L2,V,T. L1 ≛[h, o, ⓕ{I}V.T] L2 →
+                              ∧∧ L1 ≛[h, o, V] L2 & L1 ≛[h, o, T] L2.
 /2 width=2 by lfxs_inv_flat/ qed-.
 
 (* Advanced inversion lemmas ************************************************)
 
-lemma lfdeq_inv_zero_pair_sn: ∀h,o,I,Y2,L1,V1. L1.ⓑ{I}V1 ≡[h, o, #0] Y2 →
-                              ∃∃L2,V2. L1 ≡[h, o, V1] L2 & V1 ≡[h, o] V2 & Y2 = L2.ⓑ{I}V2.
-#h #o #I #Y2 #L1 #V1 #H elim (lfxs_inv_zero_pair_sn … H) -H /2 width=5 by ex3_2_intro/
-qed-.
+lemma lfdeq_inv_zero_pair_sn (h) (o): ∀I,Y2,L1,V1. L1.ⓑ{I}V1 ≛[h, o, #0] Y2 →
+                                      ∃∃L2,V2. L1 ≛[h, o, V1] L2 & V1 ≛[h, o] V2 & Y2 = L2.ⓑ{I}V2.
+/2 width=1 by lfxs_inv_zero_pair_sn/ qed-.
 
-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.
-#h #o #I #Y1 #L2 #V2 #H elim (lfxs_inv_zero_pair_dx … H) -H
-#L1 #V1 #HL12 #HV12 #H destruct /2 width=5 by ex3_2_intro/
-qed-.
+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_pair_sn: ∀h,o,I,Y2,L1,V1,i. L1.ⓑ{I}V1 ≡[h, o, #⫯i] Y2 →
-                              ∃∃L2,V2. L1 ≡[h, o, #i] L2 & Y2 = L2.ⓑ{I}V2.
-/2 width=2 by lfxs_inv_lref_pair_sn/ qed-.
+lemma lfdeq_inv_lref_bind_sn (h) (o): ∀I1,Y2,L1,i. L1.ⓘ{I1} ≛[h, o, #↑i] 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_pair_dx: ∀h,o,I,Y1,L2,V2,i. Y1 ≡[h, o, #⫯i] L2.ⓑ{I}V2 →
-                              ∃∃L1,V1. L1 ≡[h, o, #i] L2 & Y1 = L1.ⓑ{I}V1.
-/2 width=2 by lfxs_inv_lref_pair_dx/ qed-.
+lemma lfdeq_inv_lref_bind_dx (h) (o): ∀I2,Y1,L2,i. Y1 ≛[h, o, #↑i] L2.ⓘ{I2} →
+                                      ∃∃I1,L1. L1 ≛[h, o, #i] L2 & Y1 = L1.ⓘ{I1}.
+/2 width=2 by lfxs_inv_lref_bind_dx/ qed-.
 
 (* Basic forward lemmas *****************************************************)
 
+lemma lfdeq_fwd_zero_pair (h) (o): ∀I,K1,K2,V1,V2.
+                                   K1.ⓑ{I}V1 ≛[h, o, #0] K2.ⓑ{I}V2 → K1 ≛[h, o, V1] K2.
+/2 width=3 by lfxs_fwd_zero_pair/ qed-.
+
 (* Basic_2A1: uses: lleq_fwd_bind_sn lleq_fwd_flat_sn *)
-lemma lfdeq_fwd_pair_sn: ∀h,o,I,L1,L2,V,T. L1 ≡[h, o, ②{I}V.T] L2 → L1 ≡[h, o, V] L2.
+lemma lfdeq_fwd_pair_sn (h) (o): ∀I,L1,L2,V,T. L1 ≛[h, o, ②{I}V.T] L2 → L1 ≛[h, o, V] L2.
 /2 width=3 by lfxs_fwd_pair_sn/ qed-.
 
 (* Basic_2A1: uses: lleq_fwd_bind_dx lleq_fwd_bind_O_dx *)
-lemma lfdeq_fwd_bind_dx: ∀h,o,p,I,L1,L2,V,T.
-                         L1 ≡[h, o, ⓑ{p,I}V.T] L2 → L1.ⓑ{I}V ≡[h, o, T] L2.ⓑ{I}V.
+lemma lfdeq_fwd_bind_dx (h) (o): ∀p,I,L1,L2,V,T.
+                                 L1 ≛[h, o, ⓑ{p,I}V.T] L2 → L1.ⓑ{I}V ≛[h, o, T] L2.ⓑ{I}V.
 /2 width=2 by lfxs_fwd_bind_dx/ qed-.
 
 (* Basic_2A1: uses: lleq_fwd_flat_dx *)
-lemma lfdeq_fwd_flat_dx: ∀h,o,I,L1,L2,V,T. L1 ≡[h, o, ⓕ{I}V.T] L2 → L1 ≡[h, o, T] L2.
+lemma lfdeq_fwd_flat_dx (h) (o): ∀I,L1,L2,V,T. L1 ≛[h, o, ⓕ{I}V.T] L2 → L1 ≛[h, o, T] L2.
 /2 width=3 by lfxs_fwd_flat_dx/ qed-.
 
-lemma lfdeq_fwd_dx: ∀h,o,I,L1,K2,V2. ∀T:term. L1 ≡[h, o, T] K2.ⓑ{I}V2 →
-                    ∃∃K1,V1. L1 = K1.ⓑ{I}V1.
+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
-               
-*)