]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/static/lsubf_frees.ma
improved lsubf allowes to prove lsubf_frees_trans.
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / static / lsubf_frees.ma
index 47fc2f75ec691cd66c02678c213a57ce4adc5b4b..884d940bc945821c5f2cb8dfdee51574f1e08238 100644 (file)
@@ -18,9 +18,8 @@ include "basic_2/static/lsubf.ma".
 
 (* Properties with context-sensitive free variables *************************)
 
-axiom lsubf_frees_trans: ∀f2,L2,T. L2 ⊢ 𝐅*⦃T⦄ ≡ f2 → ∀f,L1. ⦃L1, f⦄ ⫃𝐅* ⦃L2, f2⦄ →
+lemma lsubf_frees_trans: ∀f2,L2,T. L2 ⊢ 𝐅*⦃T⦄ ≡ f2 → ∀f,L1. ⦃L1, f⦄ ⫃𝐅* ⦃L2, f2⦄ →
                          ∃∃f1. L1 ⊢ 𝐅*⦃T⦄ ≡ f1 & f1 ⊆ f.
-(*
 #f2 #L2 #T #H elim H -f2 -L2 -T
 [ #f2 #I #Hf2 #f #L1 #H elim (lsubf_inv_atom2 … H) -H
   #H #_ destruct /3 width=3 by frees_atom, sle_isid_sn, ex2_intro/
@@ -32,11 +31,10 @@ axiom lsubf_frees_trans: ∀f2,L2,T. L2 ⊢ 𝐅*⦃T⦄ ≡ f2 → ∀f,L1. ⦃
   [ #K1 #H elim (sle_inv_nx … H ??) -H [ <tl_next_rew |*: // ]
     #g2 #_ #H1 #H12 #H2 destruct elim (IH … H12) -K2
     /3 width=7 by frees_zero, sle_next, ex2_intro/
-  | #g #K1 #V #Hg <tl_next_rew #Hf lapply (sor_sym … Hf) -Hf
-    #Hf #H elim (sle_inv_nx … H ??) -H [|*: // ]
+  | #g #K1 #V #Hg #Hf #H elim (sle_inv_nx … H ??) -H [ <tl_next_rew |*: // ]
     #g2 #_ #H1 #H12 #H2 #H3 destruct elim (IH … H12) -K2
     #f1 #Hf1 elim (sor_isfin_ex … f1 g ??)
-    /5 width=10 by frees_fwd_isfin, frees_flat, frees_zero, monotonic_sle_sor, sor_inv_sle_dx, sor_sym, sor_sle_sn, sle_next, ex2_intro/
+    /4 width=7 by frees_fwd_isfin, frees_flat, frees_zero, sor_inv_sle, sle_next, ex2_intro/
   ]
 | #f2 #I #K2 #W #i #_ #IH #f #L1 #H elim (lsubf_inv_pair2 … H) -H *
   [ #K1 #_ #H12 #H | #g #K1 #V #Hg #Hf #_ #H12 #H1 #H2 ]
@@ -47,5 +45,14 @@ axiom lsubf_frees_trans: ∀f2,L2,T. L2 ⊢ 𝐅*⦃T⦄ ≡ f2 → ∀f,L1. ⦃
   destruct elim (IH … H12) -K2
   /3 width=3 by frees_gref, sle_inv_tl_dx, ex2_intro/
 | #f2V #f2T #f2 #p #I #L2 #V #T #_ #_ #Hf2 #IHV #IHT #f #L1 #H12
+  elim (IHV f L1) -IHV [2: /3 width=4 by lsubf_sle_div, sor_inv_sle_sn/ ]
+  elim (IHT (⫯f) (L1.ⓑ{I}V)) -IHT [2: /4 width=4 by lsubf_sle_div, lsubf_pair_nn, sor_inv_sle_dx, sor_inv_tl_dx/ ]
+  -f2V -f2T -f2 -L2 #f1T #HT #Hf1T #f1V #HV #Hf1V elim (sor_isfin_ex … f1V (⫱f1T) ??)
+  /4 width=9 by frees_fwd_isfin, frees_bind, sor_inv_sle, sle_xn_tl, isfin_tl, ex2_intro/
 | #f2V #f2T #f2 #I #L2 #V #T #_ #_ #Hf2 #IHV #IHT #f #L1 #H12
-*)
+  elim (IHV f L1) -IHV [2: /3 width=4 by lsubf_sle_div, sor_inv_sle_sn/ ]
+  elim (IHT f L1) -IHT [2: /3 width=4 by lsubf_sle_div, sor_inv_sle_dx/ ]
+  -f2V -f2T -f2 -L2 #f1T #HT #Hf1T #f1V #HV #Hf1V elim (sor_isfin_ex … f1V f1T ??)
+  /3 width=7 by frees_fwd_isfin, frees_flat, sor_inv_sle, ex2_intro/
+]
+qed-.