]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/etc_new/llpx_sn/llpx_sn_drop.etc
lfpx_drops completed!
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / etc_new / llpx_sn / llpx_sn_drop.etc
index 02e70f4fdb753df56eed7ff9cc192e286eafbcec..9bb605e39bae2cf52811f5fd70feda5b839a2f21 100644 (file)
@@ -167,246 +167,6 @@ lemma llpx_sn_dec: ∀R. (∀L,T1,T2. Decidable (R L T1 T2)) →
 -x /4 width=4 by llpx_sn_fwd_length, or_intror/
 qed-.
 
-(* Properties on relocation *************************************************)
-
-lemma llpx_sn_lift_le: ∀R. d_liftable R →
-                       ∀K1,K2,T,l0. llpx_sn R l0 T K1 K2 →
-                       ∀L1,L2,l,k. ⬇[Ⓕ, l, k] L1 ≡ K1 → ⬇[Ⓕ, l, k] L2 ≡ K2 →
-                       ∀U. ⬆[l, k] T ≡ U → l0 ≤ l → llpx_sn R l0 U L1 L2.
-#R #HR #K1 #K2 #T #l0 #H elim H -K1 -K2 -T -l0
-[ #K1 #K2 #l0 #s #HK12 #L1 #L2 #l #k #HLK1 #HLK2 #X #H #_ >(lift_inv_sort1 … H) -X
-  lapply (drop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2 -l
-  /2 width=1 by llpx_sn_sort/
-| #K1 #K2 #l0 #i #HK12 #Hil0 #L1 #L2 #l #k #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_lref1 … H) -H
-  * #Hli #H destruct
-  [ lapply (drop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2 -l
-    /2 width=1 by llpx_sn_skip/
-  | elim (ylt_yle_false … Hil0) -L1 -L2 -K1 -K2 -k -Hil0
-    /3 width=3 by yle_trans, yle_inj/
-  ]
-| #I #K1 #K2 #K11 #K22 #V1 #V2 #l0 #i #Hil0 #HK11 #HK22 #HK12 #HV12 #IHK12 #L1 #L2 #l #k #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_lref1 … H) -H
-  * #Hli #H destruct [ -HK12 | -IHK12 ]
-  [ elim (drop_trans_lt … HLK1 … HK11) // -K1
-    elim (drop_trans_lt … HLK2 … HK22) // -Hli -K2
-    /3 width=18 by llpx_sn_lref/
-  | lapply (drop_trans_ge_comm … HLK1 … HK11 ?) // -K1
-    lapply (drop_trans_ge_comm … HLK2 … HK22 ?) // -Hli -Hl0 -K2
-    /3 width=9 by llpx_sn_lref, yle_plus_dx1_trans/
-  ]
-| #K1 #K2 #l0 #i #HK1 #HK2 #HK12 #L1 #L2 #l #k #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_lref1 … H) -H
-  * #Hil #H destruct
-  lapply (drop_fwd_length_eq2 … HLK1 HLK2 HK12) -HK12
-  [ /3 width=7 by llpx_sn_free, drop_fwd_be/
-  | lapply (drop_fwd_length … HLK1) -HLK1 #HLK1
-    lapply (drop_fwd_length … HLK2) -HLK2 #HLK2
-    @llpx_sn_free [ >HLK1 | >HLK2 ] -Hil -HLK1 -HLK2 /2 width=1 by monotonic_le_plus_r/ (**) (* explicit constructor *)
-  ]
-| #K1 #K2 #l0 #p #HK12 #L1 #L2 #l #k #HLK1 #HLK2 #X #H #_ >(lift_inv_gref1 … H) -X
-  lapply (drop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2 -l -k
-  /2 width=1 by llpx_sn_gref/
-| #a #I #K1 #K2 #V #T #l0 #_ #_ #IHV #IHT #L1 #L2 #l #k #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_bind1 … H) -H
-  #W #U #HVW #HTU #H destruct /4 width=6 by llpx_sn_bind, drop_skip, yle_succ/
-| #I #K1 #K2 #V #T #l0 #_ #_ #IHV #IHT #L1 #L2 #l #k #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_flat1 … H) -H
-  #W #U #HVW #HTU #H destruct /3 width=6 by llpx_sn_flat/
-]
-qed-.
-
-lemma llpx_sn_lift_ge: ∀R,K1,K2,T,l0. llpx_sn R l0 T K1 K2 →
-                       ∀L1,L2,l,k. ⬇[Ⓕ, l, k] L1 ≡ K1 → ⬇[Ⓕ, l, k] L2 ≡ K2 →
-                       ∀U. ⬆[l, k] T ≡ U → l ≤ l0 → llpx_sn R (l0+k) U L1 L2.
-#R #K1 #K2 #T #l0 #H elim H -K1 -K2 -T -l0
-[ #K1 #K2 #l0 #s #HK12 #L1 #L2 #l #k #HLK1 #HLK2 #X #H #_ >(lift_inv_sort1 … H) -X
-  lapply (drop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2 -l
-  /2 width=1 by llpx_sn_sort/
-| #K1 #K2 #l0 #i #HK12 #Hil0 #L1 #L2 #l #k #HLK1 #HLK2 #X #H #_ elim (lift_inv_lref1 … H) -H
-  * #_ #H destruct
-  lapply (drop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2
-  [ /3 width=3 by llpx_sn_skip, ylt_plus_dx2_trans/
-  | /3 width=3 by llpx_sn_skip, monotonic_ylt_plus_dx/
-  ]
-| #I #K1 #K2 #K11 #K22 #V1 #V2 #l0 #i #Hil0 #HK11 #HK22 #HK12 #HV12 #_ #L1 #L2 #l #k #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_lref1 … H) -H
-  * #Hil #H destruct
-  [ elim (ylt_yle_false … Hil0) -I -L1 -L2 -K1 -K2 -K11 -K22 -V1 -V2 -k -Hil0
-    /3 width=3 by ylt_yle_trans, ylt_inj/
-  | lapply (drop_trans_ge_comm … HLK1 … HK11 ?) // -K1
-    lapply (drop_trans_ge_comm … HLK2 … HK22 ?) // -Hil -Hl0 -K2
-    /3 width=9 by llpx_sn_lref, monotonic_yle_plus_dx/
-  ]
-| #K1 #K2 #l0 #i #HK1 #HK2 #HK12 #L1 #L2 #l #k #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_lref1 … H) -H
-  * #Hil #H destruct
-  lapply (drop_fwd_length_eq2 … HLK1 HLK2 HK12) -HK12
-  [ /3 width=7 by llpx_sn_free, drop_fwd_be/
-  | lapply (drop_fwd_length … HLK1) -HLK1 #HLK1
-    lapply (drop_fwd_length … HLK2) -HLK2 #HLK2
-    @llpx_sn_free [ >HLK1 | >HLK2 ] -Hil -HLK1 -HLK2 /2 width=1 by monotonic_le_plus_r/ (**) (* explicit constructor *)
-  ]
-| #K1 #K2 #l0 #p #HK12 #L1 #L2 #l #k #HLK1 #HLK2 #X #H #_ >(lift_inv_gref1 … H) -X
-  lapply (drop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2 -l
-  /2 width=1 by llpx_sn_gref/
-| #a #I #K1 #K2 #V #T #l0 #_ #_ #IHV #IHT #L1 #L2 #l #k #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_bind1 … H) -H
-  #W #U #HVW #HTU #H destruct /4 width=5 by llpx_sn_bind, drop_skip, yle_succ/
-| #I #K1 #K2 #V #T #l0 #_ #_ #IHV #IHT #L1 #L2 #l #k #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_flat1 … H) -H
-  #W #U #HVW #HTU #H destruct /3 width=5 by llpx_sn_flat/
-]
-qed-.
-
-(* Inversion lemmas on relocation *******************************************)
-
-lemma llpx_sn_inv_lift_le: ∀R. d_deliftable_sn R →
-                           ∀L1,L2,U,l0. llpx_sn R l0 U L1 L2 →
-                           ∀K1,K2,l,k. ⬇[Ⓕ, l, k] L1 ≡ K1 → ⬇[Ⓕ, l, k] L2 ≡ K2 →
-                           ∀T. ⬆[l, k] T ≡ U → l0 ≤ l → llpx_sn R l0 T K1 K2.
-#R #HR #L1 #L2 #U #l0 #H elim H -L1 -L2 -U -l0
-[ #L1 #L2 #l0 #s #HL12 #K1 #K2 #l #k #HLK1 #HLK2 #X #H #_ >(lift_inv_sort2 … H) -X
-  lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -l -k
-  /2 width=1 by llpx_sn_sort/
-| #L1 #L2 #l0 #i #HL12 #Hil0 #K1 #K2 #l #k #HLK1 #HLK2 #X #H #_ elim (lift_inv_lref2 … H) -H
-  * #_ #H destruct
-  lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2
-  [ /2 width=1 by llpx_sn_skip/
-  | /3 width=3 by llpx_sn_skip, yle_ylt_trans/
-  ]
-| #I #L1 #L2 #K11 #K22 #W1 #W2 #l0 #i #Hil0 #HLK11 #HLK22 #HK12 #HW12 #IHK12 #K1 #K2 #l #k #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_lref2 … H) -H
-  * #Hil #H destruct [ -HK12 | -IHK12 ]
-  [ elim (drop_conf_lt … HLK1 … HLK11) // -L1 #L1 #V1 #HKL1 #HKL11 #HVW1
-    elim (drop_conf_lt … HLK2 … HLK22) // -Hil -L2 #L2 #V2 #HKL2 #HKL22 #HVW2
-    elim (HR … HW12 … HKL11 … HVW1) -HR #V0 #HV0 #HV12
-    lapply (lift_inj … HV0 … HVW2) -HV0 -HVW2 #H destruct
-    /3 width=10 by llpx_sn_lref/
-  | lapply (drop_conf_ge … HLK1 … HLK11 ?) // -L1
-    lapply (drop_conf_ge … HLK2 … HLK22 ?) // -L2 -Hil0
-    elim (yle_inv_plus_inj2 … Hil) -Hil /4 width=9 by llpx_sn_lref, yle_trans, yle_inj/ (**) (* slow *)
-  ]
-| #L1 #L2 #l0 #i #HL1 #HL2 #HL12 #K1 #K2 #l #k #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_lref2 … H) -H
-  * #_ #H destruct
-  lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12)
-  [ lapply (drop_fwd_length_le4 … HLK1) -HLK1
-    lapply (drop_fwd_length_le4 … HLK2) -HLK2
-    #HKL2 #HKL1 #HK12 @llpx_sn_free // /2 width=3 by transitive_le/ (**) (* full auto too slow *)
-  | lapply (drop_fwd_length … HLK1) -HLK1 #H >H in HL1; -H
-    lapply (drop_fwd_length … HLK2) -HLK2 #H >H in HL2; -H
-    /3 width=1 by llpx_sn_free, le_plus_to_minus_r/
-  ]
-| #L1 #L2 #l0 #p #HL12 #K1 #K2 #l #k #HLK1 #HLK2 #X #H #_ >(lift_inv_gref2 … H) -X
-  lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -l -k
-  /2 width=1 by llpx_sn_gref/
-| #a #I #L1 #L2 #W #U #l0 #_ #_ #IHW #IHU #K1 #K2 #l #k #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_bind2 … H) -H
-  #V #T #HVW #HTU #H destruct /4 width=6 by llpx_sn_bind, drop_skip, yle_succ/
-| #I #L1 #L2 #W #U #l0 #_ #_ #IHW #IHU #K1 #K2 #l #k #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_flat2 … H) -H
-  #V #T #HVW #HTU #H destruct /3 width=6 by llpx_sn_flat/
-]
-qed-.
-
-lemma llpx_sn_inv_lift_be: ∀R,L1,L2,U,l0. llpx_sn R l0 U L1 L2 →
-                           ∀K1,K2,l,k. ⬇[Ⓕ, l, k] L1 ≡ K1 → ⬇[Ⓕ, l, k] L2 ≡ K2 →
-                           ∀T. ⬆[l, k] T ≡ U → l ≤ l0 → l0 ≤ l + k → llpx_sn R l T K1 K2.
-#R #L1 #L2 #U #l0 #H elim H -L1 -L2 -U -l0
-[ #L1 #L2 #l0 #s #HL12 #K1 #K2 #l #k #HLK1 #HLK2 #X #H #_ #_ >(lift_inv_sort2 … H) -X
-  lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -l0 -k
-  /2 width=1 by llpx_sn_sort/
-| #L1 #L2 #l0 #i #HL12 #Hil0 #K1 #K2 #l #k #HLK1 #HLK2 #X #H #Hl0 #Hl0k elim (lift_inv_lref2 … H) -H
-  * #Hil #H destruct
-  [ lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2
-    -Hil0 /3 width=1 by llpx_sn_skip, ylt_inj/
-  | elim (ylt_yle_false … Hil0) -L1 -L2 -Hl0 -Hil0
-    /3 width=3 by yle_trans, yle_inj/ (**) (* slow *)
-  ]
-| #I #L1 #L2 #K11 #K22 #W1 #W2 #l0 #i #Hil0 #HLK11 #HLK22 #HK12 #HW12 #_ #K1 #K2 #l #k #HLK1 #HLK2 #X #H #Hl0 #Hl0k elim (lift_inv_lref2 … H) -H
-  * #Hil #H destruct
-  [ elim (ylt_yle_false … Hil0) -I -L1 -L2 -K11 -K22 -W1 -W2 -Hl0k -Hil0
-    /3 width=3 by ylt_yle_trans, ylt_inj/
-  | lapply (drop_conf_ge … HLK1 … HLK11 ?) // -L1
-    lapply (drop_conf_ge … HLK2 … HLK22 ?) // -L2 -Hil0 -Hl0 -Hl0k
-    elim (yle_inv_plus_inj2 … Hil) -Hil /3 width=9 by llpx_sn_lref, yle_inj/
-  ]
-| #L1 #L2 #l0 #i #HL1 #HL2 #HL12 #K1 #K2 #l #k #HLK1 #HLK2 #X #H #Hl0 #Hl0k elim (lift_inv_lref2 … H) -H
-  * #_ #H destruct
-  lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12)
-  [ lapply (drop_fwd_length_le4 … HLK1) -HLK1
-    lapply (drop_fwd_length_le4 … HLK2) -HLK2
-    #HKL2 #HKL1 #HK12 @llpx_sn_free // /2 width=3 by transitive_le/ (**) (* full auto too slow *)
-  | lapply (drop_fwd_length … HLK1) -HLK1 #H >H in HL1; -H
-    lapply (drop_fwd_length … HLK2) -HLK2 #H >H in HL2; -H
-    /3 width=1 by llpx_sn_free, le_plus_to_minus_r/
-  ]
-| #L1 #L2 #l0 #p #HL12 #K1 #K2 #l #k #HLK1 #HLK2 #X #H #_ #_ >(lift_inv_gref2 … H) -X
-  lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -l0 -k
-  /2 width=1 by llpx_sn_gref/
-| #a #I #L1 #L2 #W #U #l0 #_ #_ #IHW #IHU #K1 #K2 #l #k #HLK1 #HLK2 #X #H #Hl0 #Hl0k elim (lift_inv_bind2 … H) -H
-  #V #T #HVW #HTU #H destruct
-  @llpx_sn_bind [ /2 width=5 by/ ] -IHW (**) (* explicit constructor *)
-  @(IHU … HTU) -IHU -HTU /2 width=1 by drop_skip, yle_succ/
-| #I #L1 #L2 #W #U #l0 #_ #_ #IHW #IHU #K1 #K2 #l #k #HLK1 #HLK2 #X #H #Hl0 #Hl0k elim (lift_inv_flat2 … H) -H
-  #V #T #HVW #HTU #H destruct /3 width=6 by llpx_sn_flat/
-]
-qed-.
-
-lemma llpx_sn_inv_lift_ge: ∀R,L1,L2,U,l0. llpx_sn R l0 U L1 L2 →
-                           ∀K1,K2,l,k. ⬇[Ⓕ, l, k] L1 ≡ K1 → ⬇[Ⓕ, l, k] L2 ≡ K2 →
-                           ∀T. ⬆[l, k] T ≡ U → l + k ≤ l0 → llpx_sn R (l0-k) T K1 K2.
-#R #L1 #L2 #U #l0 #H elim H -L1 -L2 -U -l0
-[ #L1 #L2 #l0 #s #HL12 #K1 #K2 #l #k #HLK1 #HLK2 #X #H #_ >(lift_inv_sort2 … H) -X
-  lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -l
-  /2 width=1 by llpx_sn_sort/
-| #L1 #L2 #l0 #i #HL12 #Hil0 #K1 #K2 #l #k #HLK1 #HLK2 #X #H #Hlml0 elim (lift_inv_lref2 … H) -H
-  * #Hil #H destruct [ -Hil0 | -Hlml0 ]
-  lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2
-  [ /4 width=3 by llpx_sn_skip, yle_plus1_to_minus_inj2, ylt_yle_trans, ylt_inj/
-  | elim (yle_inv_plus_inj2 … Hil) -Hil
-    /3 width=1 by llpx_sn_skip, monotonic_ylt_minus_dx/
-  ]
-| #I #L1 #L2 #K11 #K22 #W1 #W2 #l0 #i #Hil0 #HLK11 #HLK22 #HK12 #HW12 #_ #K1 #K2 #l #k #HLK1 #HLK2 #X #H #Hlml0 elim (lift_inv_lref2 … H) -H
-  * #Hil #H destruct
-  [ elim (ylt_yle_false … Hil0) -I -L1 -L2 -K11 -K22 -W1 -W2 -Hil0
-    /3 width=3 by yle_fwd_plus_sn1, ylt_yle_trans, ylt_inj/
-  | lapply (drop_conf_ge … HLK1 … HLK11 ?) // -L1
-    lapply (drop_conf_ge … HLK2 … HLK22 ?) // -L2 -Hlml0 -Hil
-    /3 width=9 by llpx_sn_lref, monotonic_yle_minus_dx/
-  ]
-| #L1 #L2 #l0 #i #HL1 #HL2 #HL12 #K1 #K2 #l #k #HLK1 #HLK2 #X #H #Hlml0 elim (lift_inv_lref2 … H) -H
-  * #_ #H destruct
-  lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12)
-  [ lapply (drop_fwd_length_le4 … HLK1) -HLK1
-    lapply (drop_fwd_length_le4 … HLK2) -HLK2
-    #HKL2 #HKL1 #HK12 @llpx_sn_free // /2 width=3 by transitive_le/ (**) (* full auto too slow *)
-  | lapply (drop_fwd_length … HLK1) -HLK1 #H >H in HL1; -H
-    lapply (drop_fwd_length … HLK2) -HLK2 #H >H in HL2; -H
-    /3 width=1 by llpx_sn_free, le_plus_to_minus_r/
-  ]
-| #L1 #L2 #l0 #p #HL12 #K1 #K2 #l #k #HLK1 #HLK2 #X #H #_ >(lift_inv_gref2 … H) -X
-  lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -l
-  /2 width=1 by llpx_sn_gref/
-| #a #I #L1 #L2 #W #U #l0 #_ #_ #IHW #IHU #K1 #K2 #l #k #HLK1 #HLK2 #X #H #Hlml0 elim (lift_inv_bind2 … H) -H
-  #V #T #HVW #HTU #H destruct
-  @llpx_sn_bind [ /2 width=5 by/ ] -IHW (**) (* explicit constructor *)
-  <yminus_succ1_inj /2 width=2 by yle_fwd_plus_sn2/
-  @(IHU … HTU) -IHU -HTU /2 width=1 by drop_skip, yle_succ/
-| #I #L1 #L2 #W #U #l0 #_ #_ #IHW #IHU #K1 #K2 #l #k #HLK1 #HLK2 #X #H #Hlml0 elim (lift_inv_flat2 … H) -H
-  #V #T #HVW #HTU #H destruct /3 width=5 by llpx_sn_flat/
-]
-qed-.
-
-(* Advanced inversion lemmas on relocation **********************************)
-
-lemma llpx_sn_inv_lift_O: ∀R,L1,L2,U. llpx_sn R 0 U L1 L2 →
-                          ∀K1,K2,k. ⬇[k] L1 ≡ K1 → ⬇[k] L2 ≡ K2 →
-                          ∀T. ⬆[0, k] T ≡ U → llpx_sn R 0 T K1 K2.
-/2 width=11 by llpx_sn_inv_lift_be/ qed-.
-
-lemma llpx_sn_drop_conf_O: ∀R,L1,L2,U. llpx_sn R 0 U L1 L2 →
-                           ∀K1,k. ⬇[k] L1 ≡ K1 → ∀T. ⬆[0, k] T ≡ U →
-                           ∃∃K2. ⬇[k] L2 ≡ K2 & llpx_sn R 0 T K1 K2.
-#R #L1 #L2 #U #HU #K1 #k #HLK1 #T #HTU elim (llpx_sn_fwd_drop_sn … HU … HLK1)
-/3 width=10 by llpx_sn_inv_lift_O, ex2_intro/
-qed-.
-
-lemma llpx_sn_drop_trans_O: ∀R,L1,L2,U. llpx_sn R 0 U L1 L2 →
-                            ∀K2,k. ⬇[k] L2 ≡ K2 → ∀T. ⬆[0, k] T ≡ U →
-                            ∃∃K1. ⬇[k] L1 ≡ K1 & llpx_sn R 0 T K1 K2.
-#R #L1 #L2 #U #HU #K2 #k #HLK2 #T #HTU elim (llpx_sn_fwd_drop_dx … HU … HLK2)
-/3 width=10 by llpx_sn_inv_lift_O, ex2_intro/
-qed-.
-
 (* Inversion lemmas on negated lazy pointwise extension *********************)
 
 lemma nllpx_sn_inv_bind: ∀R. (∀L,T1,T2. Decidable (R L T1 T2)) →