]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/relocation/drops.ma
- new component "s_transition" for the restored fqu and fquq
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / relocation / drops.ma
index 1fc24f51e56e63b193c619a26c8c9436c8164131..2e21e55ef35f87d1541f969ebd78f317be44448c 100644 (file)
@@ -13,6 +13,7 @@
 (**************************************************************************)
 
 include "ground_2/relocation/rtmap_isfin.ma".
+include "basic_2/notation/relations/rdropstar_3.ma".
 include "basic_2/notation/relations/rdropstar_4.ma".
 include "basic_2/relocation/lreq.ma".
 include "basic_2/relocation/lifts.ma".
@@ -31,6 +32,9 @@ inductive drops (c:bool): rtmap → relation lenv ≝
               drops c (↑f) (L1.ⓑ{I}V1) (L2.ⓑ{I}V2)
 .
 
+interpretation "uniform slicing (local environment)"
+   'RDropStar i L1 L2 = (drops true (uni i) L1 L2).
+
 interpretation "general slicing (local environment)"
    'RDropStar c f L1 L2 = (drops c f L1 L2).
 
@@ -64,69 +68,6 @@ definition dedropable_sn: predicate (rtmap → relation lenv) ≝
                           ∀f2. f ⊚ f1 ≡ f2 →
                           ∃∃L2. R f2 L1 L2 & ⬇*[c, f] L2 ≡ K2 & L1 ≡[f] L2.
 
-(* Basic inversion lemmas ***************************************************)
-
-fact drops_inv_atom1_aux: ∀X,Y,c,f. ⬇*[c, f] X ≡ Y → X = ⋆ →
-                          Y = ⋆ ∧ (c = Ⓣ → 𝐈⦃f⦄).
-#X #Y #c #f * -X -Y -f
-[ /3 width=1 by conj/
-| #I #L1 #L2 #V #f #_ #H destruct
-| #I #L1 #L2 #V1 #V2 #f #_ #_ #H destruct
-]
-qed-.
-
-(* Basic_1: includes: drop_gen_sort *)
-(* Basic_2A1: includes: drop_inv_atom1 *)
-lemma drops_inv_atom1: ∀Y,c,f. ⬇*[c, f] ⋆ ≡ Y → Y = ⋆ ∧ (c = Ⓣ → 𝐈⦃f⦄).
-/2 width=3 by drops_inv_atom1_aux/ qed-.
-
-fact drops_inv_drop1_aux: ∀X,Y,c,f. ⬇*[c, f] X ≡ Y → ∀I,K,V,g. X = K.ⓑ{I}V → f = ⫯g →
-                          ⬇*[c, g] K ≡ Y.
-#X #Y #c #f * -X -Y -f
-[ #f #Ht #J #K #W #g #H destruct
-| #I #L1 #L2 #V #f #HL #J #K #W #g #H1 #H2 <(injective_next … H2) -g destruct //
-| #I #L1 #L2 #V1 #V2 #f #_ #_ #J #K #W #g #_ #H2 elim (discr_push_next … H2)
-]
-qed-.
-
-(* Basic_1: includes: drop_gen_drop *)
-(* Basic_2A1: includes: drop_inv_drop1_lt drop_inv_drop1 *)
-lemma drops_inv_drop1: ∀I,K,Y,V,c,f. ⬇*[c, ⫯f] K.ⓑ{I}V ≡ Y → ⬇*[c, f] K ≡ Y.
-/2 width=7 by drops_inv_drop1_aux/ qed-.
-
-
-fact drops_inv_skip1_aux: ∀X,Y,c,f. ⬇*[c, f] X ≡ Y → ∀I,K1,V1,g. X = K1.ⓑ{I}V1 → f = ↑g →
-                          ∃∃K2,V2. ⬇*[c, g] K1 ≡ K2 & ⬆*[g] V2 ≡ V1 & Y = K2.ⓑ{I}V2.
-#X #Y #c #f * -X -Y -f
-[ #f #Ht #J #K1 #W1 #g #H destruct
-| #I #L1 #L2 #V #f #_ #J #K1 #W1 #g #_ #H2 elim (discr_next_push … H2)
-| #I #L1 #L2 #V1 #V2 #f #HL #HV #J #K1 #W1 #g #H1 #H2 <(injective_push … H2) -g destruct
-  /2 width=5 by ex3_2_intro/
-]
-qed-.
-
-(* Basic_1: includes: drop_gen_skip_l *)
-(* Basic_2A1: includes: drop_inv_skip1 *)
-lemma drops_inv_skip1: ∀I,K1,V1,Y,c,f. ⬇*[c, ↑f] K1.ⓑ{I}V1 ≡ Y →
-                       ∃∃K2,V2. ⬇*[c, f] K1 ≡ K2 & ⬆*[f] V2 ≡ V1 & Y = K2.ⓑ{I}V2.
-/2 width=5 by drops_inv_skip1_aux/ qed-.
-
-fact drops_inv_skip2_aux: ∀X,Y,c,f. ⬇*[c, f] X ≡ Y → ∀I,K2,V2,g. Y = K2.ⓑ{I}V2 → f = ↑g →
-                          ∃∃K1,V1. ⬇*[c, g] K1 ≡ K2 & ⬆*[g] V2 ≡ V1 & X = K1.ⓑ{I}V1.
-#X #Y #c #f * -X -Y -f
-[ #f #Ht #J #K2 #W2 #g #H destruct
-| #I #L1 #L2 #V #f #_ #J #K2 #W2 #g #_ #H2 elim (discr_next_push … H2)
-| #I #L1 #L2 #V1 #V2 #f #HL #HV #J #K2 #W2 #g #H1 #H2 <(injective_push … H2) -g destruct
-  /2 width=5 by ex3_2_intro/
-]
-qed-.
-
-(* Basic_1: includes: drop_gen_skip_r *)
-(* Basic_2A1: includes: drop_inv_skip2 *)
-lemma drops_inv_skip2: ∀I,X,K2,V2,c,f. ⬇*[c, ↑f] X ≡ K2.ⓑ{I}V2 →
-                       ∃∃K1,V1. ⬇*[c, f] K1 ≡ K2 & ⬆*[f] V2 ≡ V1 & X = K1.ⓑ{I}V1.
-/2 width=5 by drops_inv_skip2_aux/ qed-.
-
 (* Basic properties *********************************************************)
 
 lemma drops_eq_repl_back: ∀L1,L2,c. eq_repl_back … (λf. ⬇*[c, f] L1 ≡ L2).
@@ -150,6 +91,27 @@ lemma drops_refl: ∀c,L,f. 𝐈⦃f⦄ → ⬇*[c, f] L ≡ L.
 /3 width=1 by drops_skip, lifts_refl/
 qed.
 
+(* Basic_2A1: includes: drop_split *)
+lemma drops_split_trans: ∀L1,L2,f,c. ⬇*[c, f] L1 ≡ L2 → ∀f1,f2. f1 ⊚ f2 ≡ f → 𝐔⦃f1⦄ →
+                         ∃∃L. ⬇*[c, f1] L1 ≡ L & ⬇*[c, f2] L ≡ L2.
+#L1 #L2 #f #c #H elim H -L1 -L2 -f
+[ #f #Hc #f1 #f2 #Hf #Hf1 @(ex2_intro … (⋆)) @drops_atom
+  #H lapply (Hc H) -c
+  #H elim (after_inv_isid3 … Hf H) -f //
+| #I #L1 #L2 #V #f #HL12 #IHL12 #f1 #f2 #Hf #Hf1 elim (after_inv_xxn … Hf) -Hf [1,3: * |*: // ]
+  [ #g1 #g2 #Hf #H1 #H2 destruct
+    lapply (isuni_inv_push … Hf1 ??) -Hf1 [1,2: // ] #Hg1
+    elim (IHL12 … Hf) -f
+    /4 width=5 by drops_drop, drops_skip, lifts_refl, isuni_isid, ex2_intro/
+  | #g1 #Hf #H destruct elim (IHL12 … Hf) -f
+    /3 width=5 by ex2_intro, drops_drop, isuni_inv_next/
+  ]
+| #I #L1 #L2 #V1 #V2 #f #_ #HV21 #IHL12 #f1 #f2 #Hf #Hf1 elim (after_inv_xxp … Hf) -Hf [2,3: // ]
+  #g1 #g2 #Hf #H1 #H2 destruct elim (lifts_split_trans … HV21 … Hf) -HV21
+  elim (IHL12 … Hf) -f /3 width=5 by ex2_intro, drops_skip, isuni_fwd_push/
+]
+qed-.
+
 (* Basic_2A1: includes: drop_FT *)
 lemma drops_TF: ∀L1,L2,f. ⬇*[Ⓣ, f] L1 ≡ L2 → ⬇*[Ⓕ, f] L1 ≡ L2.
 #L1 #L2 #f #H elim H -L1 -L2 -f
@@ -204,17 +166,114 @@ qed-.
 lemma drops_isuni_fwd_drop2: ∀I,X,K,V,c,f. 𝐔⦃f⦄ → ⬇*[c, f] X ≡ K.ⓑ{I}V → ⬇*[c, ⫯f] X ≡ K.
 /3 width=7 by drops_after_fwd_drop2, after_isid_isuni/ qed-.
 
-(* forward lemmas with test for finite colength *****************************)
+(* Forward lemmas with test for finite colength *****************************)
 
 lemma drops_fwd_isfin: ∀L1,L2,f. ⬇*[Ⓣ, f] L1 ≡ L2 → 𝐅⦃f⦄.
 #L1 #L2 #f #H elim H -L1 -L2 -f
 /3 width=1 by isfin_next, isfin_push, isfin_isid/
 qed-.
 
-(* Basic_2A1: removed theorems 14:
+(* Basic inversion lemmas ***************************************************)
+
+fact drops_inv_atom1_aux: ∀X,Y,c,f. ⬇*[c, f] X ≡ Y → X = ⋆ →
+                          Y = ⋆ ∧ (c = Ⓣ → 𝐈⦃f⦄).
+#X #Y #c #f * -X -Y -f
+[ /3 width=1 by conj/
+| #I #L1 #L2 #V #f #_ #H destruct
+| #I #L1 #L2 #V1 #V2 #f #_ #_ #H destruct
+]
+qed-.
+
+(* Basic_1: includes: drop_gen_sort *)
+(* Basic_2A1: includes: drop_inv_atom1 *)
+lemma drops_inv_atom1: ∀Y,c,f. ⬇*[c, f] ⋆ ≡ Y → Y = ⋆ ∧ (c = Ⓣ → 𝐈⦃f⦄).
+/2 width=3 by drops_inv_atom1_aux/ qed-.
+
+fact drops_inv_drop1_aux: ∀X,Y,c,f. ⬇*[c, f] X ≡ Y → ∀I,K,V,g. X = K.ⓑ{I}V → f = ⫯g →
+                          ⬇*[c, g] K ≡ Y.
+#X #Y #c #f * -X -Y -f
+[ #f #Ht #J #K #W #g #H destruct
+| #I #L1 #L2 #V #f #HL #J #K #W #g #H1 #H2 <(injective_next … H2) -g destruct //
+| #I #L1 #L2 #V1 #V2 #f #_ #_ #J #K #W #g #_ #H2 elim (discr_push_next … H2)
+]
+qed-.
+
+(* Basic_1: includes: drop_gen_drop *)
+(* Basic_2A1: includes: drop_inv_drop1_lt drop_inv_drop1 *)
+lemma drops_inv_drop1: ∀I,K,Y,V,c,f. ⬇*[c, ⫯f] K.ⓑ{I}V ≡ Y → ⬇*[c, f] K ≡ Y.
+/2 width=7 by drops_inv_drop1_aux/ qed-.
+
+
+fact drops_inv_skip1_aux: ∀X,Y,c,f. ⬇*[c, f] X ≡ Y → ∀I,K1,V1,g. X = K1.ⓑ{I}V1 → f = ↑g →
+                          ∃∃K2,V2. ⬇*[c, g] K1 ≡ K2 & ⬆*[g] V2 ≡ V1 & Y = K2.ⓑ{I}V2.
+#X #Y #c #f * -X -Y -f
+[ #f #Ht #J #K1 #W1 #g #H destruct
+| #I #L1 #L2 #V #f #_ #J #K1 #W1 #g #_ #H2 elim (discr_next_push … H2)
+| #I #L1 #L2 #V1 #V2 #f #HL #HV #J #K1 #W1 #g #H1 #H2 <(injective_push … H2) -g destruct
+  /2 width=5 by ex3_2_intro/
+]
+qed-.
+
+(* Basic_1: includes: drop_gen_skip_l *)
+(* Basic_2A1: includes: drop_inv_skip1 *)
+lemma drops_inv_skip1: ∀I,K1,V1,Y,c,f. ⬇*[c, ↑f] K1.ⓑ{I}V1 ≡ Y →
+                       ∃∃K2,V2. ⬇*[c, f] K1 ≡ K2 & ⬆*[f] V2 ≡ V1 & Y = K2.ⓑ{I}V2.
+/2 width=5 by drops_inv_skip1_aux/ qed-.
+
+fact drops_inv_skip2_aux: ∀X,Y,c,f. ⬇*[c, f] X ≡ Y → ∀I,K2,V2,g. Y = K2.ⓑ{I}V2 → f = ↑g →
+                          ∃∃K1,V1. ⬇*[c, g] K1 ≡ K2 & ⬆*[g] V2 ≡ V1 & X = K1.ⓑ{I}V1.
+#X #Y #c #f * -X -Y -f
+[ #f #Ht #J #K2 #W2 #g #H destruct
+| #I #L1 #L2 #V #f #_ #J #K2 #W2 #g #_ #H2 elim (discr_next_push … H2)
+| #I #L1 #L2 #V1 #V2 #f #HL #HV #J #K2 #W2 #g #H1 #H2 <(injective_push … H2) -g destruct
+  /2 width=5 by ex3_2_intro/
+]
+qed-.
+
+(* Basic_1: includes: drop_gen_skip_r *)
+(* Basic_2A1: includes: drop_inv_skip2 *)
+lemma drops_inv_skip2: ∀I,X,K2,V2,c,f. ⬇*[c, ↑f] X ≡ K2.ⓑ{I}V2 →
+                       ∃∃K1,V1. ⬇*[c, f] K1 ≡ K2 & ⬆*[f] V2 ≡ V1 & X = K1.ⓑ{I}V1.
+/2 width=5 by drops_inv_skip2_aux/ qed-.
+
+(* Inversion lemmas with test for uniformity ********************************)
+
+(* Basic_2A1: was: drop_inv_O1_pair1 *)
+lemma drops_inv_pair1_isuni: ∀I,K,L2,V,c,f. 𝐔⦃f⦄ → ⬇*[c, f] K.ⓑ{I}V ≡ L2 →
+                             (𝐈⦃f⦄ ∧ L2 = K.ⓑ{I}V) ∨
+                             ∃∃g. ⬇*[c, g] K ≡ L2 & f = ⫯g.
+#I #K #L2 #V #c #f #Hf #H elim (isuni_split … Hf) -Hf * #g #Hg #H0 destruct
+[ lapply (drops_inv_skip1 … H) -H * #Y #X #HY #HX #H destruct
+  <(drops_fwd_isid … HY Hg) -Y >(lifts_fwd_isid … HX Hg) -X
+  /4 width=3 by isid_push, or_introl, conj/
+| lapply (drops_inv_drop1 … H) -H /3 width=3 by ex2_intro, or_intror/
+]
+qed-.
+
+(* Basic_2A1: was: drop_inv_O1_pair2 *)
+lemma drops_inv_pair2_isuni: ∀I,K,V,c,f,L1. 𝐔⦃f⦄ → ⬇*[c, f] L1 ≡ K.ⓑ{I}V →
+                             (𝐈⦃f⦄ ∧ L1 = K.ⓑ{I}V) ∨
+                             ∃∃I1,K1,V1,g. ⬇*[c, g] K1 ≡ K.ⓑ{I}V & L1 = K1.ⓑ{I1}V1 & f = ⫯g.
+#I #K #V #c #f *
+[ #Hf #H elim (drops_inv_atom1 … H) -H #H destruct
+| #L1 #I1 #V1 #Hf #H elim (drops_inv_pair1_isuni … Hf H) -Hf -H *
+  [ #Hf #H destruct /3 width=1 by or_introl, conj/
+  | /3 width=7 by ex3_4_intro, or_intror/
+  ]
+]
+qed-.
+
+lemma drops_inv_pair2_isuni_next: ∀I,K,V,c,f,L1. 𝐔⦃f⦄ → ⬇*[c, ⫯f] L1 ≡ K.ⓑ{I}V →
+                                  ∃∃I1,K1,V1. ⬇*[c, f] K1 ≡ K.ⓑ{I}V & L1 = K1.ⓑ{I1}V1.
+#I #K #V #c #f #L1 #Hf #H elim (drops_inv_pair2_isuni … H) -H /2 width=3 by isuni_next/ -Hf *
+[ #H elim (isid_inv_next … H) -H //
+| /2 width=5 by ex2_3_intro/
+]
+qed-. 
+
+(* Basic_2A1: removed theorems 12:
               drops_inv_nil drops_inv_cons d1_liftable_liftables
-              drop_refl_atom_O2
-              drop_inv_O1_pair1 drop_inv_pair1 drop_inv_O1_pair2
+              drop_refl_atom_O2 drop_inv_pair1
               drop_inv_Y1 drop_Y1 drop_O_Y drop_fwd_Y2
               drop_fwd_length_minus2 drop_fwd_length_minus4
 *)