]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/static/lfxs.ma
advances on lfsx ...
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / static / lfxs.ma
index a2f6816349f65e35b29f149e44d9ec32e30c52df..d7dae247e3ccbb1e041377d1936d505ae00c47e2 100644 (file)
@@ -43,7 +43,7 @@ definition R_confluent2_lfxs: relation4 (relation3 lenv term term)
                               ∀L1. L0 ⦻*[RP1, T0] L1 → ∀L2. L0 ⦻*[RP2, T0] L2 →
                               ∃∃T. R2 L1 T1 T & R1 L2 T2 T.
 
-(* Basic properties ***********************************************************)
+(* Basic properties *********************************************************)
 
 lemma lfxs_atom: ∀R,I. ⋆ ⦻*[R, ⓪{I}] ⋆.
 /3 width=3 by lexs_atom, frees_atom, ex2_intro/
@@ -89,13 +89,21 @@ lemma lfxs_co: ∀R1,R2. (∀L,T1,T2. R1 L T1 T2 → R2 L T1 T2) →
 #R1 #R2 #HR #L1 #L2 #T * /4 width=7 by lexs_co, ex2_intro/
 qed-.
 
+lemma lfxs_isid: ∀R1,R2,L1,L2,T1,T2.
+                 (∀f. L1 ⊢ 𝐅*⦃T1⦄ ≡ f → 𝐈⦃f⦄) → 
+                 (∀f. 𝐈⦃f⦄ → L1 ⊢ 𝐅*⦃T2⦄ ≡ f) → 
+                 L1 ⦻*[R1, T1] L2 → L1 ⦻*[R2, T2] L2.
+#R1 #R2 #L1 #L2 #T1 #T2 #H1 #H2 *
+/4 width=7 by lexs_co_isid, ex2_intro/
+qed-.
+
 (* Basic inversion lemmas ***************************************************)
 
-lemma lfxs_inv_atom_sn: ∀R,I,Y2. ⋆ ⦻*[R, ⓪{I}] Y2 → Y2 = ⋆.
-#R #I #Y2 * /2 width=4 by lexs_inv_atom1/
+lemma lfxs_inv_atom_sn: ∀R,Y2,T. ⋆ ⦻*[R, T] Y2 → Y2 = ⋆.
+#R #Y2 #T * /2 width=4 by lexs_inv_atom1/
 qed-.
 
-lemma lfxs_inv_atom_dx: ∀R,I,Y1. Y1 ⦻*[R, ⓪{I}] ⋆ → Y1 = ⋆.
+lemma lfxs_inv_atom_dx: ∀R,Y1,T. Y1 ⦻*[R, T] ⋆ → Y1 = ⋆.
 #R #I #Y1 * /2 width=4 by lexs_inv_atom2/
 qed-.
 
@@ -253,12 +261,20 @@ lemma lfxs_fwd_pair_sn: ∀R,I,L1,L2,V,T. L1 ⦻*[R, ②{I}V.T] L2 → L1 ⦻*[R
 #R * /2 width=4 by lfxs_fwd_flat_sn, lfxs_fwd_bind_sn/
 qed-.
 
-(* Basic_2A1: removed theorems 24:
+lemma lfxs_fwd_dx: ∀R,I,L1,K2,T,V2. L1 ⦻*[R, T] K2.ⓑ{I}V2 →
+                   ∃∃K1,V1. L1 = K1.ⓑ{I}V1.
+#R #I #L1 #K2 #T #V2 * #f elim (pn_split f) * #g #Hg #_ #Hf destruct
+[ elim (lexs_inv_push2 … Hf) | elim (lexs_inv_next2 … Hf) ] -Hf #K1 #V1 #_ #_ #H destruct
+/2 width=3 by ex1_2_intro/
+qed-.
+
+(* Basic_2A1: removed theorems 25:
               llpx_sn_sort llpx_sn_skip llpx_sn_lref llpx_sn_free llpx_sn_gref
               llpx_sn_bind llpx_sn_flat
               llpx_sn_inv_bind llpx_sn_inv_flat
               llpx_sn_fwd_lref llpx_sn_fwd_pair_sn llpx_sn_fwd_length
               llpx_sn_fwd_bind_sn llpx_sn_fwd_bind_dx llpx_sn_fwd_flat_sn llpx_sn_fwd_flat_dx
               llpx_sn_refl llpx_sn_Y llpx_sn_bind_O llpx_sn_ge_up llpx_sn_ge llpx_sn_co
-              llpx_sn_fwd_drop_sn llpx_sn_fwd_drop_dx              
+              llpx_sn_fwd_drop_sn llpx_sn_fwd_drop_dx
+              llpx_sn_dec      
 *)