]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_etc.ma
advances on cpx_lfxs_conf_fle
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_transition / cpx_etc.ma
index 7a516ee462b55ac8eb9b7d86eed95e24f25b1599..19999ed9b7953886e359107eb46367080df73829 100644 (file)
 include "basic_2/static/fle_drops.ma".
 include "basic_2/static/fle_fqup.ma".
 include "basic_2/static/fle_fle.ma".
-include "basic_2/static/lfxs.ma".
+include "basic_2/static/lfxs_length.ma".
 include "basic_2/rt_transition/cpx.ma".
 
-axiom fle_pair_sn: ∀L,V1,V2. ⦃L, V1⦄ ⊆ ⦃L, V2⦄ →
-                   ∀I1,I2,T. ⦃L.ⓑ{I1}V1, T⦄ ⊆ ⦃L.ⓑ{I2}V2, T⦄.
-(*
-axiom fle_elim_flat: ∀L1,L2,V1,T2. ⦃L1, V1⦄ ⊆ ⦃L2, T2⦄ → ∀T1. ⦃L1, T1⦄ ⊆ ⦃L2, T2⦄ →
-                     ∀I. ⦃L1, ⓕ{I}V1.T1⦄ ⊆ ⦃L2, T2⦄.
 
-axiom fle_elim_bind: 
-                     ∀p,I1,I2,J1,J2,L,V,W,T. ⦃L, ⓑ{p,I1}ⓕ{J1}V.W.T⦄ ⊆ ⦃L, ⓕ{J2}V.ⓑ{p,I2}W.T⦄.
-*)
+lemma fle_zero_bi: ∀K1,K2. |K1| = |K2| → ∀V1,V2. ⦃K1, V1⦄ ⊆ ⦃K2, V2⦄ →
+                   ∀I1,I2. ⦃K1.ⓑ{I1}V1, #O⦄ ⊆ ⦃K2.ⓑ{I2}V2, #O⦄.
+#K1 #K2 #HK #V1 #V2
+* #n1 #n2 #f1 #f2 #Hf1 #Hf2 #HK12 #Hf12
+#I1 #I2
+elim (lveq_inj_length … HK12) // -HK #H1 #H2 destruct
+/3 width=12 by frees_pair, lveq_bind, sle_next, ex4_4_intro/
+qed.
+
 (* UNCOUNTED CONTEXT-SENSITIVE PARALLEL RT-TRANSITION FOR TERMS *************)
 
 (* Properties with context-sensitive free variables *************************)
 
 (* Note: "⦃L2, T1⦄ ⊆ ⦃L0, T1⦄" may not hold *)
-axiom cpx_lfxs_conf_gen: ∀R. c_reflexive … R →
-                         ∀h,G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ⬈[h] T1 →
+axiom cpx_lfxs_conf_fle: ∀R,h. c_reflexive … R →
+                         (∨∨ (∀G. (cpx h G) = R) | R_fle_compatible R) →
+                         ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ⬈[h] T1 →
                          ∀L2. L0 ⪤*[R, T0] L2 →
                          ∧∧ ⦃L2, T0⦄ ⊆ ⦃L0, T0⦄ & ⦃L2, T1⦄ ⊆ ⦃L2, T0⦄
                           & ⦃L0, T1⦄ ⊆ ⦃L0, T0⦄.
 (*
-#R #HR #h #G #L0 #T0 @(fqup_wf_ind_eq (Ⓕ) … G L0 T0) -G -L0 -T0
+#R #h #H1R #H2R #G #L0 #T0 @(fqup_wf_ind_eq (Ⓕ) … G L0 T0) -G -L0 -T0
 #G #L #T #IH #G0 #L0 * *
-[ #s #HG #HL #HT #X #HX #Y #_ destruct -IH
+[ #s #HG #HL #HT #X #HX #Y #HY destruct -IH
+  lapply (lfxs_fwd_length … HY) -HY #H0
   elim (cpx_inv_sort1 … HX) -HX #H destruct
-  /2 width=1 by and3_intro/
-|
-| #l #HG #HL #HT #X #HX #Y #_ destruct -IH
+  /3 width=1 by fle_sort_length, and3_intro/
+| * [| #i ] #HG #HL #HT #X #HX #Y #HY destruct
+  [ elim (cpx_inv_zero1 … HX) -HX
+    [ #H destruct
+      elim (lfxs_inv_zero … HY) -HY *
+      [ #H1 #H2 destruct -IH /2 width=1 by and3_intro/
+      | #I #K0 #K2 #V0 #V2 #HK02 #HV02 #H1 #H2 destruct
+        lapply (lfxs_fwd_length … HK02) #HK
+        elim H2R -H2R #H2R
+        [ <(H2R G0) in HV02; -H2R #HV02
+          elim (IH … HV02 … HK02) /2 width=2 by fqu_fqup, fqu_lref_O/ -IH -HV02 -HK02 #H1V #H2V #_
+          /4 width=1 by fle_trans_tc, fle_zero_bi, and3_intro/
+        | lapply (H2R … HV02 … HK02) -H2R -HV02 -HK02 -IH #HKV20
+          /3 width=1 by fle_zero_bi, and3_intro/
+        ]
+      | #f #I #K0 #K2 #Hf #HK02 #H1 #H2 destruct
+      ]
+    | * #I0 #K0 #V0 #V1 #HV01 #HV1X #H destruct
+      elim (lfxs_inv_zero_pair_sn … HY) -HY #K2 #V2 #HK02 #HV02 #H destruct
+    ]
+  | elim (cpx_inv_lref1 … HX) -HX
+    [ #H destruct
+      elim (lfxs_inv_lref … HY) -HY *
+      [ #H0 #H1 destruct /2 width=1 by and3_intro/
+      | #I0 #I2 #K0 #K2 #HK02 #H1 #H2 destruct
+        lapply (lfxs_fwd_length … HK02) #HK
+        elim (IH … HK02) [|*: /2 width=4 by fqu_fqup/ ] -IH -HK02
+        /3 width=5 by and3_intro, fle_lifts_SO/
+      ]
+    | * #I0 #K0 #V1 #HV1 #HV1X #H0 destruct
+      elim (lfxs_inv_lref_bind_sn … HY) -HY #I2 #K2 #HK02 #H destruct
+      lapply (lfxs_fwd_length … HK02) #HK
+      elim (IH … HK02) [|*: /2 width=4 by fqu_fqup/ ] -IH -HV1 -HK02
+      /3 width=5 by fle_lifts_SO, and3_intro/
+    ]
+  ]
+| #l #HG #HL #HT #X #HX #Y #HY destruct -IH
+  lapply (lfxs_fwd_length … HY) -HY #H0
   >(cpx_inv_gref1 … HX) -X
-  /2 width=1 by and3_intro/
+  /3 width=1 by fle_gref_length, and3_intro/
 | #p #I #V0 #T0 #HG #HL #HT #X #HX #Y #HY destruct
+  lapply (lfxs_fwd_length … HY) #H0
   elim (lfxs_inv_bind … V0 ? HY) -HY // #HV0 #HT0
   elim (cpx_inv_bind1 … HX) -HX *
   [ #V1 #T1 #HV01 #HT01 #H destruct
     elim (IH … HV01 … HV0) -HV01 -HV0 // #H1V #H2V #H3V
     elim (IH … HT01 … HT0) -HT01 -HT0 -IH // #H1T #H2T #H3T
-    /5 width=4 by fle_pair_sn, fle_trans, fle_bind, and3_intro/
+    /4 width=3 by fle_bind_eq, fle_fwd_pair_sn, and3_intro/
   | #T #HT #HXT #H1 #H2 destruct
     elim (IH G0 … V0… V0 … HV0) -HV0 // #H1V #H2V #H3V
     elim (IH … HT … HT0) -HT -HT0 -IH // #H1T #H2T #H3T
-    /4 width=7 by fle_bind_dx, fle_trans, fle_bind, and3_intro/
+    /3 width=5 by fle_bind, fle_inv_lifts_sn, and3_intro/
   ]
 | #I #V0 #X0 #HG #HL #HT #X #HX #Y #HY destruct
   elim (lfxs_inv_flat … HY) -HY #HV0 #HX0
@@ -69,25 +108,32 @@ axiom cpx_lfxs_conf_gen: ∀R. c_reflexive … R →
   | #HX #H destruct
     elim (IH G0 … V0… V0 … HV0) -HV0 // #H1V #H2V #H3V
     elim (IH … HX … HX0) -HX -HX0 -IH // #H1T #H2T #H3T
-    /4 width=4 by fle_trans, fle_flat, and3_intro/
+    /4 width=3 by fle_flat_sn, fle_flat_dx_dx, fle_flat_dx_sn, and3_intro/
   | #HX #H destruct
     elim (IH … HX … HV0) -HX -HV0 // #H1V #H2V #H3V
     elim (IH G0 … X0… X0 … HX0) -HX0 -IH // #H1T #H2T #H3T
-    /4 width=4 by fle_trans, fle_flat, and3_intro/
+    /4 width=3 by fle_flat_sn, fle_flat_dx_dx, fle_flat_dx_sn, and3_intro/
   | #p #V1 #W0 #W1 #T0 #T1 #HV01 #HW01 #HT01 #H1 #H2 #H3 destruct
     elim (lfxs_inv_bind … W0 ? HX0) -HX0 // #HW0 #HT0
     elim (IH … HV01 … HV0) -HV01 -HV0 // #H1V #H2V #H3V
     elim (IH … HW01 … HW0) -HW01 -HW0 // #H1W #H2W #H3W
     elim (IH … HT01 … HT0) -HT01 -HT0 -IH // #H1T #H2T #H3T
-    @and3_intro
-    [ /3 width=5 by fle_flat, fle_bind/
-    | @(fle_trans)
-      [4: @(fle_flat … H2V) [2: @(fle_bind … H2T) // | skip | @Appl ]
-      |1,2: skip
-      | @(fle_trans) [3: @fle_elim_bind  
-    
-     … H2T) // 
-    
-    // |1,2: skip
-    | /2 width=1 by fle_bind_dx/      
+    lapply (fle_fwd_pair_sn … H2T) -H2T #H2T
+    lapply (fle_fwd_pair_sn … H3T) -H3T #H3T
+    @and3_intro [ /3 width=5 by fle_flat, fle_bind/ ] (**) (* full auto too slow *)
+    @fle_bind_sn_ge /4 width=1 by fle_shift, fle_flat_sn, fle_flat_dx_dx, fle_flat_dx_sn, fle_bind_dx_sn/
+  | #p #V1 #X1 #W0 #W1 #T0 #T1 #HV01 #HVX1 #HW01 #HT01 #H1 #H2 #H3 destruct
+    elim (lfxs_inv_bind … W0 ? HX0) -HX0 // #HW0 #HT0
+    elim (IH … HV01 … HV0) -HV01 -HV0 // #H1V #H2V #H3V
+    elim (IH … HW01 … HW0) -HW01 -HW0 // #H1W #H2W #H3W
+    elim (IH … HT01 … HT0) -HT01 -HT0 -IH // #H1T #H2T #H3T
+    lapply (fle_fwd_pair_sn … H2T) -H2T #H2T
+    lapply (fle_fwd_pair_sn … H3T) -H3T #H3T
+    @and3_intro[ /3 width=5 by fle_flat, fle_bind/ ] (**) (* full auto too slow *)
+    @fle_bind_sn_ge //
+    [1,3: /3 width=1 by fle_flat_dx_dx, fle_bind_dx_sn/
+    |2,4: /4 width=3 by fle_flat_sn, fle_flat_dx_sn, fle_flat_dx_dx, fle_shift, fle_lifts_sn/
+    ]
+  ]
+]
 *)