]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/relocation/ldrop_leq.ma
- advances on hereditarily free variables: now "frees" is primitive
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / relocation / ldrop_leq.ma
index 179bb5f050d541cb0a5dc26fbb6e1f8792b7cd0e..75586d3d3b5b27ae8b826af9cc898fb6ffb09709 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-include "ground_2/ynat/ynat_plus.ma".
-include "basic_2/grammar/leq.ma".
+include "basic_2/grammar/leq_leq.ma".
 include "basic_2/relocation/ldrop.ma".
 
 (* BASIC SLICING FOR LOCAL ENVIRONMENTS *************************************)
 
-lemma ldrop_leq_conf_ge: ∀L1,L2,d,e. L1 ≃[d, e] L2 →
-                         ∀I,K,V,i. ⇩[O, i]L1 ≡ K.ⓑ{I}V → d + e ≤ i →
-                         ⇩[O, i]L2 ≡ K.ⓑ{I}V.
+definition dedropable_sn: predicate (relation lenv) ≝
+                          λR. ∀L1,K1,s,d,e. ⇩[s, d, e] L1 ≡ K1 → ∀K2. R K1 K2 →
+                          ∃∃L2. R L1 L2 & ⇩[s, d, e] L2 ≡ K2 & L1 ≃[d, e] L2.
+
+(* Properties on equivalence ************************************************)
+
+lemma leq_ldrop_trans_be: ∀L1,L2,d,e. L1 ≃[d, e] L2 →
+                          ∀I,K2,W,s,i. ⇩[s, 0, i] L2 ≡ K2.ⓑ{I}W →
+                          d ≤ i → i < d + e →
+                          ∃∃K1. K1 ≃[0, ⫰(d+e-i)] K2 & ⇩[s, 0, i] L1 ≡ K1.ⓑ{I}W.
 #L1 #L2 #d #e #H elim H -L1 -L2 -d -e
-[ #d #e #J #K #W #i #H elim (ldrop_inv_atom1 … H) -H
-  #H destruct
-| #I #L1 #L2 #V #HL12 #IHL12 #J #K #W #i #H #_ elim (ldrop_inv_O1_pair1 … H) -H
-  * #H1 #H2
-  [ -IHL12 lapply (leq_inv_O2 … HL12) -HL12
-    #H3 destruct //
-  | -HL12 /4 width=1 by ldrop_ldrop_lt, yle_inj/
+[ #d #e #J #K2 #W #s #i #H
+  elim (ldrop_inv_atom1 … H) -H #H destruct
+| #I1 #I2 #L1 #L2 #V1 #V2 #_ #_ #J #K2 #W #s #i #_ #_ #H
+  elim (ylt_yle_false … H) //
+| #I #L1 #L2 #V #e #HL12 #IHL12 #J #K2 #W #s #i #H #_ >yplus_O1
+  elim (ldrop_inv_O1_pair1 … H) -H * #Hi #HLK1 [ -IHL12 | -HL12 ]
+  [ #_ destruct >ypred_succ
+    /2 width=3 by ldrop_pair, ex2_intro/
+  | lapply (ylt_inv_O1 i ?) /2 width=1 by ylt_inj/
+    #H <H -H #H lapply (ylt_inv_succ … H) -H
+    #Hie elim (IHL12 … HLK1) -IHL12 -HLK1 // -Hie
+    >yminus_succ <yminus_inj /3 width=3 by ldrop_drop_lt, ex2_intro/
   ]
-| #I1 #I2 #L1 #L2 #V1 #V2 #e #_ #IHL12 #J #K #W #i #H1 >yplus_succ_swap
-  #Hei elim (yle_inv_inj2 … Hei) -Hei
-  #x #Hei #H elim (yplus_inv_inj … H) -H normalize
-  #y #z >commutative_plus
-  #H1 #H2 #H3 destruct elim (le_inv_plus_l … Hei) -Hei
-  #Hzi #Hi lapply (ldrop_inv_ldrop1_lt … H1 ?) -H1
-  /4 width=1 by ldrop_ldrop_lt, yle_inj/
-| #I #L1 #L2 #V #d #e #_ #IHL12 #J #K #W #i #H0 #Hdei elim (yle_inv_inj2 … Hdei) -Hdei
-  #x #Hdei #H elim (yplus_inv_inj … H) -H
-  #y #z >commutative_plus
-  #H1 #H2 #H3 destruct elim (ysucc_inv_inj_dx … H2) -H2
-  #x #H1 #H2 destruct elim (le_inv_plus_l … Hdei)
-  #_ #Hi lapply (ldrop_inv_ldrop1_lt … H0 ?) -H0
-  [2: #H0 @ldrop_ldrop_lt ] [2,3: /2 width=3 by lt_to_le_to_lt/ ]
-  /4 width=3 by yle_inj, monotonic_pred/
+| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #IHL12 #J #K2 #W #s #i #HLK2 #Hdi
+  elim (yle_inv_succ1 … Hdi) -Hdi
+  #Hdi #Hi <Hi >yplus_succ1 #H lapply (ylt_inv_succ … H) -H
+  #Hide lapply (ldrop_inv_drop1_lt … HLK2 ?) -HLK2 /2 width=1 by ylt_O/
+  #HLK1 elim (IHL12 … HLK1) -IHL12 -HLK1 <yminus_inj >yminus_SO2
+  /4 width=3 by ylt_O, ldrop_drop_lt, ex2_intro/
 ]
 qed-.
 
-lemma ldrop_leq_conf_be: ∀L1,L2,d,e. L1 ≃[d, e] L2 →
-                         ∀I1,K1,V1,i. ⇩[O, i]L1 ≡ K1.ⓑ{I1}V1 → d ≤ i → i < d + e →
-                         ∃∃I2,K2,V2. K1 ≃[0, ⫰(d+e-i)] K2 & ⇩[O, i]L2 ≡ K2.ⓑ{I2}V2.
-#L1 #L2 #d #e #H elim H -L1 -L2 -d -e
-[ #d #e #J1 #K1 #W1 #i #H elim (ldrop_inv_atom1 … H) -H
-  #H destruct
-| #I #L1 #L2 #V #HL12 #IHL12 #J1 #K1 #W1 #i #_ #_ #H elim (ylt_yle_false … H) //
-| #I1 #I2 #L1 #L2 #V1 #V2 #e #HL12 >yplus_O1 >yplus_O1
-  #IHL12 #J1 #K1 #W1 #i #H #_ elim (eq_or_gt i) #Hi destruct [ -IHL12 | -HL12 ]
-  [ #_ lapply (ldrop_inv_O2 … H) -H
-    #H destruct >ypred_succ /2 width=5 by ldrop_pair, ex2_3_intro/
-  | lapply (ldrop_inv_ldrop1_lt … H ?) -H //
-    #H <(ylt_inv_O1 i) /2 width=1 by ylt_inj/
-    #Hie lapply (ylt_inv_succ … Hie) -Hie
-    #Hie elim (IHL12 … H) -IHL12 -H //
-    >yminus_succ /3 width=5 by ldrop_ldrop_lt, ex2_3_intro/
-  ]
-| #I #L1 #L2 #V #d #e #_ #IHL12 #J1 #K1 #W1 #i #H #Hdi lapply (ylt_yle_trans 0 … Hdi ?) //
-  #Hi <(ylt_inv_O1 … Hi) >yplus_succ1 >yminus_succ elim (yle_inv_succ1 … Hdi) -Hdi
-  #Hdi #_ #Hide lapply (ylt_inv_succ … Hide)
-  #Hide lapply (ylt_inv_inj … Hi) -Hi
-  #Hi lapply (ldrop_inv_ldrop1_lt … H ?) -H //
-  #H elim (IHL12 … H) -IHL12 -H
-  /3 width=5 by ldrop_ldrop_lt, ex2_3_intro/
-] 
+lemma leq_ldrop_conf_be: ∀L1,L2,d,e. L1 ≃[d, e] L2 →
+                         ∀I,K1,W,s,i. ⇩[s, 0, i] L1 ≡ K1.ⓑ{I}W →
+                         d ≤ i → i < d + e →
+                         ∃∃K2. K1 ≃[0, ⫰(d+e-i)] K2 & ⇩[s, 0, i] L2 ≡ K2.ⓑ{I}W.
+#L1 #L2 #d #e #HL12 #I #K1 #W #s #i #HLK1 #Hdi #Hide
+elim (leq_ldrop_trans_be … (leq_sym … HL12) … HLK1) // -L1 -Hdi -Hide
+/3 width=3 by leq_sym, ex2_intro/
 qed-.
 
-lemma ldrop_leq_conf_lt: ∀L1,L2,d,e. L1 ≃[d, e] L2 →
-                         ∀I,K1,V,i. ⇩[O, i]L1 ≡ K1.ⓑ{I}V → i < d →
-                         ∃∃K2. K1 ≃[⫰(d-i), e] K2 & ⇩[O, i]L2 ≡ K2.ⓑ{I}V.
-#L1 #L2 #d #e #H elim H -L1 -L2 -d -e
-[ #d #e #J #K1 #W #i #H elim (ldrop_inv_atom1 … H) -H
-  #H destruct
-| #I #L1 #L2 #V #_ #_ #J #K1 #W #i #_ #H elim (ylt_yle_false … H) //
-| #I1 #I2 #L1 #L2 #V1 #V2 #e #_ #_ #J #K1 #W #i #_ #H elim (ylt_yle_false … H) //
-| #I #L1 #L2 #V #d #e #HL12 #IHL12 #J #K1 #W #i #H elim (eq_or_gt i) #Hi destruct [ -IHL12 | -HL12 ]
-  [ #_ lapply (ldrop_inv_O2 … H) -H
-    #H destruct >ypred_succ /2 width=5 by ldrop_pair, ex2_intro/
-  | lapply (ldrop_inv_ldrop1_lt … H ?) -H //
-    #H <(ylt_inv_O1 i) /2 width=1 by ylt_inj/
-    #Hie lapply (ylt_inv_succ … Hie) -Hie
-    #Hie elim (IHL12 … H) -IHL12 -H
-    >yminus_succ /3 width=5 by ldrop_ldrop_lt, ex2_intro/
-  ]
+lemma ldrop_O1_ex: ∀K2,i,L1. |L1| = |K2| + i →
+                   ∃∃L2. L1 ≃[0, i] L2 & ⇩[i] L2 ≡ K2.
+#K2 #i @(nat_ind_plus … i) -i
+[ /3 width=3 by leq_O2, ex2_intro/
+| #i #IHi #Y #Hi elim (ldrop_O1_lt (Ⓕ) Y 0) //
+  #I #L1 #V #H lapply (ldrop_inv_O2 … H) -H #H destruct
+  normalize in Hi; elim (IHi L1) -IHi
+  /3 width=5 by ldrop_drop, leq_pair, injective_plus_l, ex2_intro/
+]
+qed-.
+
+lemma dedropable_sn_TC: ∀R. dedropable_sn R → dedropable_sn (TC … R).
+#R #HR #L1 #K1 #s #d #e #HLK1 #K2 #H elim H -K2
+[ #K2 #HK12 elim (HR … HLK1 … HK12) -HR -K1
+  /3 width=4 by inj, ex3_intro/
+| #K #K2 #_ #HK2 * #L #H1L1 #HLK #H2L1 elim (HR … HLK … HK2) -HR -K
+  /3 width=6 by leq_trans, step, ex3_intro/
+]
+qed-.
+
+(* Inversion lemmas on equivalence ******************************************)
+
+lemma ldrop_O1_inj: ∀i,L1,L2,K. ⇩[i] L1 ≡ K → ⇩[i] L2 ≡ K → L1 ≃[i, ∞] L2.
+#i @(nat_ind_plus … i) -i
+[ #L1 #L2 #K #H <(ldrop_inv_O2 … H) -K #H <(ldrop_inv_O2 … H) -L1 //
+| #i #IHi * [2: #L1 #I1 #V1 ] * [2,4: #L2 #I2 #V2 ] #K #HLK1 #HLK2 //
+  lapply (ldrop_fwd_length … HLK1)
+  <(ldrop_fwd_length … HLK2) [ /4 width=5 by ldrop_inv_drop1, leq_succ/ ]
+  normalize <plus_n_Sm #H destruct
 ]
 qed-.