]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/reduction/lpx_frees.ma
- advances on free variables allow to reduce lleq_lpx_trans to llor_total :)
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / reduction / lpx_frees.ma
index 1458cf3fa69c898c40e48c902309baa2bc04a089..b44d88504da2202c005c532786c312c292e29f69 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/multiple/fqup.ma".
+include "basic_2/multiple/frees_leq.ma".
 include "basic_2/multiple/frees_lift.ma".
 include "basic_2/reduction/lpx_ldrop.ma".
 
 (* SN EXTENDED PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS ********************)
-(*
-lemma cofrees_lsuby_conf: āˆ€L1,U,i. L1 āŠ¢ i ~Ļµ š…*ā¦ƒUā¦„ ā†’
-                          āˆ€L2. lsuby L1 L2 ā†’ L2 āŠ¢ i ~Ļµ š…*ā¦ƒUā¦„.
-/3 width=3 by lsuby_cpys_trans/ qed-.
 
-lemma lpx_cpx_cofrees_conf: āˆ€h,g,G,L1,U1,U2. ā¦ƒG, L1ā¦„ āŠ¢ U1 āž”[h, g] U2 ā†’
-                            āˆ€L2. ā¦ƒG, L1ā¦„ āŠ¢ āž”[h, g] L2 ā†’
-                            āˆ€i. L1 āŠ¢ i ~Ļµ š…*ā¦ƒU1ā¦„ ā†’ L2 āŠ¢ i ~Ļµ š…*ā¦ƒU2ā¦„.
+(* Properties on context-sensitive free variables ***************************)
+
+lemma lpx_cpx_frees_trans: āˆ€h,g,G,L1,U1,U2. ā¦ƒG, L1ā¦„ āŠ¢ U1 āž”[h, g] U2 ā†’
+                           āˆ€L2. ā¦ƒG, L1ā¦„ āŠ¢ āž”[h, g] L2 ā†’
+                           āˆ€i. L2 āŠ¢ i Ļµ š…*[0]ā¦ƒU2ā¦„ ā†’ L1 āŠ¢ i Ļµ š…*[0]ā¦ƒU1ā¦„.
 #h #g #G #L1 #U1 @(fqup_wf_ind_eq ā€¦ G L1 U1) -G -L1 -U1
 #G0 #L0 #U0 #IH #G #L1 * *
-[ -IH #k #HG #HL #HU #U2 #H elim (cpx_inv_sort1 ā€¦ H) -H
-  [| * #l #_ ] #H destruct //
-| #j #HG #HL #HU #U2 #H #L2 #HL12 #i #Hi elim (cpx_inv_lref1 ā€¦ H) -H
-  [ #H destruct elim (lt_or_eq_or_gt i j) #Hji
-    [ -IH -HL12 /2 width=4 by cofrees_lref_gt/
-    | -IH -HL12 destruct elim (cofrees_inv_lref_eq ā€¦ Hi)
-    | elim (lt_or_ge j (|L2|)) /2 width=5 by cofrees_lref_free/ #Hj
-      elim (ldrop_O1_lt (ā’») L1 j) [2: >(lpx_fwd_length ā€¦ HL12) // ] #I #K1 #W1 #HLK1
-      elim (lpx_ldrop_conf ā€¦ HLK1 ā€¦ HL12) -HL12 #Y #H #HLK2
-      elim (lpx_inv_pair1 ā€¦ H) -H #K2 #W2 #HK12 #HW12 #H destruct
-      lapply (cofrees_inv_lref_lt ā€¦ Hi ā€¦ HLK1) -Hi // #HW1
-      lapply (IH ā€¦ HW12 ā€¦ HK12 ā€¦ HW1) /2 width=2 by fqup_lref/ -L1 -K1 -W1 #HW2
-      /2 width=9 by cofrees_lref_lt/ (**) (* full auto too slow *)
-    ]
-  | * #I #K1 #W1 #W0 #HLK1 #HW10 #HW0U2 elim (lt_or_ge j i) #Hji
-    [ lapply (ldrop_fwd_drop2 ā€¦ HLK1) #H0
-      elim (lpx_ldrop_conf ā€¦ H0 ā€¦ HL12) #K2 #HK12 #HLK2
-      @(cofrees_lift_ge ā€¦ HLK2 ā€¦ HW0U2) //
-      @(IH ā€¦ HW10) /2 width=2 by fqup_lref/ -L2 -K2 -W0 -U2 -IH
-      <minus_plus /2 width=7 by cofrees_inv_lref_lt/
-    | -IH lapply (ldrop_fwd_drop2 ā€¦ HLK1) -HLK1 #HLK1
-      elim (lpx_ldrop_conf ā€¦ HLK1 ā€¦ HL12) -I -L1 -W1
-      /2 width=12 by cofrees_lift_be/
+[ -IH #k #HG #HL #HU #U2 #H1 #L2 #_ #i #H2 elim (cpx_inv_sort1 ā€¦ H1) -H1
+  [| * #l #_ ] #H destruct elim (frees_inv_sort ā€¦ H2)
+| #j #HG #HL #HU #U2 #H1 #L2 #HL12 #i #H2 elim (cpx_inv_lref1 ā€¦ H1) -H1
+  [ #H destruct elim (frees_inv_lref ā€¦ H2) -H2 //
+    * #I #K2 #W2 #Hj #Hji #HLK2 #HW2
+    elim (lpx_ldrop_trans_O1 ā€¦ HL12 ā€¦ HLK2) -HL12 #Y #HLK1 #H
+    elim (lpx_inv_pair2 ā€¦ H) -H #K1 #W1 #HK12 #HW12 #H destruct
+    /4 width=11 by frees_lref_be, fqup_lref/
+  | * #I #K1 #W1 #W0 #HLK1 #HW10 #HW0U2
+    lapply (ldrop_fwd_drop2 ā€¦ HLK1) #H0
+    elim (lpx_ldrop_conf ā€¦ H0 ā€¦ HL12) -H0 -HL12 #K2 #HK12 #HLK2
+    elim (lt_or_ge i (j+1)) #Hji
+    [ -IH elim (frees_inv_lift_be ā€¦ H2 ā€¦ HLK2 ā€¦ HW0U2) /2 width=1 by monotonic_pred/
+    | lapply (frees_inv_lift_ge ā€¦ H2 ā€¦ HLK2 ā€¦ HW0U2 ?) -L2 -U2 // <minus_plus destruct
+      /4 width=11 by frees_lref_be, fqup_lref/
     ]
   ]
-| -IH #p #HG #HL #HU #U2 #H lapply (cpx_inv_gref1 ā€¦ H) -H
-  #H destruct //
+| -IH #p #HG #HL #HU #U2 #H1 >(cpx_inv_gref1 ā€¦ H1) -H1 destruct
+   #L2 #_ #i #H2 elim (frees_inv_gref ā€¦ H2)
 | #a #I #W1 #U1 #HG #HL #HU #X #HX #L2 #HL12 #i #Hi destruct
-  elim (cofrees_inv_bind ā€¦ Hi) -Hi #HW1 #HU1
-  elim (cpx_inv_bind1 ā€¦ HX) -HX
-  [ * #W2 #U2 #HW12 #HU12 #H destruct /4 width=8 by cofrees_bind, lpx_pair/
-  |
+  elim (cpx_inv_bind1 ā€¦ HX) -HX *
+  [ #W2 #U2 #HW12 #HU12 #H destruct
+    elim (frees_inv_bind_O ā€¦ Hi) -Hi
+    /4 width=7 by frees_bind_dx_O, frees_bind_sn, lpx_pair/
+  | #U2 #HU12 #HXU2 #H1 #H2 destruct
+    lapply (frees_lift_ge ā€¦ Hi (L2.ā““W1) (ā’») ā€¦ HXU2 ?)
+    /4 width=7 by frees_bind_dx_O, lpx_pair, ldrop_drop/
   ]
 | #I #W1 #X1 #HG #HL #HU #X2 #HX2 #L2 #HL12 #i #Hi destruct
-  elim (cofrees_inv_flat ā€¦ Hi) -Hi #HW1 #HX1
   elim (cpx_inv_flat1 ā€¦ HX2) -HX2 *
-  [ #W2 #U2 #HW12 #HU12 #H destruct /3 width=7 by cofrees_flat/
-  | #HU12 #H destruct /2 width=7 by/
-  | #HW12 #H destruct /2 width=7 by/
+  [ #W2 #U2 #HW12 #HU12 #H destruct
+    elim (frees_inv_flat ā€¦ Hi) -Hi /3 width=7 by frees_flat_dx, frees_flat_sn/
+  | #HU12 #H destruct /3 width=7 by frees_flat_dx/
+  | #HW12 #H destruct /3 width=7 by frees_flat_sn/
   | #b #W2 #V1 #V2 #U1 #U2 #HW12 #HV12 #HU12 #H1 #H2 #H3 destruct
-    elim (cofrees_inv_bind ā€¦ HX1) -HX1 #HV1 #HU1
-    @cofrees_bind [ /3 width=7 by cofrees_flat/ ]
-    @(cofrees_lsuby_conf (L2.ā“›V2)) /3 width=7 by lpx_pair, lsuby_succ/
-  | 
-    
-   
-*)
+    elim (frees_inv_bind ā€¦ Hi) -Hi #Hi
+    [ elim (frees_inv_flat ā€¦ Hi) -Hi
+      /4 width=7 by frees_flat_dx, frees_flat_sn, frees_bind_sn/
+    | lapply (leq_frees_trans ā€¦ Hi (L2.ā“›V2) ?) /2 width=1 by leq_succ/ -Hi #HU2
+      lapply (frees_weak ā€¦ HU2 0 ?) -HU2
+      /5 width=7 by frees_bind_dx_O, frees_flat_dx, lpx_pair/
+    ]
+  | #b #W2 #W0 #V1 #V2 #U1 #U2 #HW12 #HW20 #HV12 #HU12 #H1 #H2 #H3 destruct
+    elim (frees_inv_bind_O ā€¦ Hi) -Hi #Hi
+    [ /4 width=7 by frees_flat_dx, frees_bind_sn/
+    | elim (frees_inv_flat ā€¦ Hi) -Hi
+      [ #HW0 lapply (frees_inv_lift_ge ā€¦ HW0 L2 (ā’») ā€¦ HW20 ?) -W0
+        /3 width=7 by frees_flat_sn, ldrop_drop/
+      | /5 width=7 by frees_bind_dx_O, frees_flat_dx, lpx_pair/
+      ]
+    ]
+  ]
+]
+qed-.
+
+lemma cpx_frees_trans: āˆ€h,g,G. frees_trans (cpx h g G).
+/2 width=8 by lpx_cpx_frees_trans/ qed-.
+
+lemma lpx_frees_trans: āˆ€h,g,G,L1,L2. ā¦ƒG, L1ā¦„ āŠ¢ āž”[h, g] L2 ā†’
+                       āˆ€U,i. L2 āŠ¢ i Ļµ š…*[0]ā¦ƒUā¦„ ā†’ L1 āŠ¢ i Ļµ š…*[0]ā¦ƒUā¦„.
+/2 width=8 by lpx_cpx_frees_trans/ qed-.